Skip to content

[ refactor ] Data.Fin.Properties of decidable equality, plus knock-ons #2740

New issue

Have a question about this project? Sign up for a free GitHub account to open an issue and contact its maintainers and the community.

By clicking “Sign up for GitHub”, you agree to our terms of service and privacy statement. We’ll occasionally send you account related emails.

Already on GitHub? Sign in to your account

Open
wants to merge 19 commits into
base: master
Choose a base branch
from
Open
Show file tree
Hide file tree
Changes from all commits
Commits
File filter

Filter by extension

Filter by extension

Conversations
Failed to load comments.
Loading
Jump to
Jump to file
Failed to load files.
Loading
Diff view
Diff view
19 changes: 18 additions & 1 deletion CHANGELOG.md
Original file line number Diff line number Diff line change
Expand Up @@ -251,9 +251,21 @@ Additions to existing modules
swap : Permutation m n → Permutation (suc (suc m)) (suc (suc n))
```

* In `Data.Fin.Permutation.Components`:
```agda
transpose-iij : (i j : Fin n) → transpose i i j ≡ j
transpose-ijj : (i j : Fin n) → transpose i j j ≡ i
transpose-iji : (i j : Fin n) → transpose i j i ≡ j
transpose-transpose : transpose i j k ≡ l → transpose j i l ≡ k
```

* In `Data.Fin.Properties`:
```agda
cast-involutive : .(eq₁ : m ≡ n) .(eq₂ : n ≡ m) → ∀ k → cast eq₁ (cast eq₂ k) ≡ k
≡-irrelevant : Irrelevant {A = Fin n} _≡_
≟-diag : (eq : i ≡ j) → (i ≟ j) ≡ yes eq
≟-diag-refl : (i : Fin n) → (i ≟ i) ≡ yes refl
≟-off-diag : (i≢j : i ≢ j) → (i ≟ j) ≡ no i≢j
```

* In `Data.Fin.Subset`:
Expand All @@ -262,7 +274,6 @@ Additions to existing modules
_⊉_ : Subset n → Subset n → Set
_⊃_ : Subset n → Subset n → Set
_⊅_ : Subset n → Subset n → Set

```

* In `Data.Fin.Subset.Induction`:
Expand Down Expand Up @@ -337,6 +348,12 @@ Additions to existing modules
map-id : map id ≗ id {A = List⁺ A}
```

* In `Data.Nat.Properties`:
```agda
≟-diag-refl : ∀ n → (n ≟ n) ≡ yes refl
≟-off-diag : (m≢n : m ≢ n) → (m ≟ n) ≡ no m≢n
```

* In `Data.Product.Function.Dependent.Propositional`:
```agda
Σ-↪ : (I↪J : I ↪ J) → (∀ {j} → A (from I↪J j) ↪ B j) → Σ I A ↪ Σ J B
Expand Down
68 changes: 32 additions & 36 deletions src/Data/Fin/Permutation.agda
Original file line number Diff line number Diff line change
Expand Up @@ -11,12 +11,14 @@ module Data.Fin.Permutation where
open import Data.Bool.Base using (true; false)
open import Data.Fin.Base using (Fin; suc; cast; opposite; punchIn; punchOut)
open import Data.Fin.Patterns using (0F; 1F)
open import Data.Fin.Properties using (punchInᵢ≢i; punchOut-punchIn;
punchOut-cong; punchOut-cong′; punchIn-punchOut
; _≟_; ¬Fin0; cast-involutive; opposite-involutive)
open import Data.Fin.Properties
using (¬Fin0; _≟_; ≟-diag-refl; ≟-off-diag
; cast-involutive; opposite-involutive
; punchInᵢ≢i; punchOut-punchIn; punchIn-punchOut
; punchOut-cong; punchOut-cong′)
import Data.Fin.Permutation.Components as PC
open import Data.Nat.Base using (ℕ; suc; zero)
open import Data.Product.Base using (_,_; proj₂)
open import Data.Product.Base using (_,_)
open import Function.Bundles using (_↔_; Injection; Inverse; mk↔ₛ′)
open import Function.Construct.Composition using (_↔-∘_)
open import Function.Construct.Identity using (↔-id)
Expand All @@ -26,9 +28,8 @@ open import Function.Properties.Inverse using (↔⇒↣)
open import Function.Base using (_∘_; _∘′_)
open import Level using (0ℓ)
open import Relation.Binary.Core using (Rel)
open import Relation.Nullary using (does; ¬_; yes; no)
open import Relation.Nullary.Decidable using (dec-yes; dec-no)
open import Relation.Nullary.Negation using (contradiction)
open import Relation.Nullary.Decidable.Core using (does; yes; no)
open import Relation.Nullary.Negation.Core using (¬_; contradiction)
open import Relation.Binary.PropositionalEquality.Core
using (_≡_; _≢_; refl; sym; trans; subst; cong; cong₂)
open import Relation.Binary.PropositionalEquality.Properties
Expand Down Expand Up @@ -167,19 +168,19 @@ remove {m} {n} i π = permutation to from inverseˡ′ inverseʳ′

inverseʳ′ : StrictlyInverseʳ _≡_ to from
inverseʳ′ j = begin
from (to j) ≡⟨⟩
punchOut {i = i} {πˡ (punchIn (πʳ i) (punchOut to-punchOut))} _ ≡⟨ punchOut-cong′ i (cong πˡ (punchIn-punchOut _)) ⟩
punchOut {i = i} {πˡ (πʳ (punchIn i j))} _ ≡⟨ punchOut-cong i (inverseˡ π) ⟩
punchOut {i = i} {punchIn i j} _ ≡⟨ punchOut-punchIn i ⟩
j
from (to j) ≡⟨⟩
punchOut {i = i} {πˡ (punchIn (πʳ i) (punchOut to-punchOut))} _ ≡⟨ punchOut-cong′ i (cong πˡ (punchIn-punchOut _)) ⟩
punchOut {i = i} {πˡ (πʳ (punchIn i j))} _ ≡⟨ punchOut-cong i (inverseˡ π) ⟩
punchOut {i = i} {punchIn i j} _ ≡⟨ punchOut-punchIn i ⟩
j ∎

inverseˡ′ : StrictlyInverseˡ _≡_ to from
inverseˡ′ j = begin
to (from j) ≡⟨⟩
punchOut {i = πʳ i} {πʳ (punchIn i (punchOut from-punchOut))} _ ≡⟨ punchOut-cong′ (πʳ i) (cong πʳ (punchIn-punchOut _)) ⟩
punchOut {i = πʳ i} {πʳ (πˡ (punchIn (πʳ i) j))} _ ≡⟨ punchOut-cong (πʳ i) (inverseʳ π) ⟩
punchOut {i = πʳ i} {punchIn (πʳ i) j} _ ≡⟨ punchOut-punchIn (πʳ i) ⟩
j
to (from j) ≡⟨⟩
punchOut {i = πʳ i} {πʳ (punchIn i (punchOut from-punchOut))} _ ≡⟨ punchOut-cong′ (πʳ i) (cong πʳ (punchIn-punchOut _)) ⟩
punchOut {i = πʳ i} {πʳ (πˡ (punchIn (πʳ i) j))} _ ≡⟨ punchOut-cong (πʳ i) (inverseʳ π) ⟩
punchOut {i = πʳ i} {punchIn (πʳ i) j} _ ≡⟨ punchOut-punchIn (πʳ i) ⟩
j ∎

------------------------------------------------------------------------
-- Lifting
Expand Down Expand Up @@ -228,23 +229,23 @@ insert {m} {n} i j π = permutation to from inverseˡ′ inverseʳ′

inverseʳ′ : StrictlyInverseʳ _≡_ to from
inverseʳ′ k with i ≟ k
... | yes i≡k rewrite proj₂ (dec-yes (j ≟ j) refl) = i≡k
... | yes i≡k rewrite ≟-diag-refl j = i≡k
... | no i≢k
with j≢punchInⱼπʳpunchOuti≢k ← punchInᵢ≢i j (π ⟨$⟩ʳ punchOut i≢k) ∘ sym
rewrite dec-no (j ≟ punchIn j (π ⟨$⟩ʳ punchOut i≢k)) j≢punchInⱼπʳpunchOuti≢k
rewrite ≟-off-diag j≢punchInⱼπʳpunchOuti≢k
= begin
punchIn i (π ⟨$⟩ˡ punchOut j≢punchInⱼπʳpunchOuti≢k) ≡⟨ cong (λ l → punchIn i (π ⟨$⟩ˡ l)) (punchOut-cong j refl) ⟩
punchIn i (π ⟨$⟩ˡ punchOut (punchInᵢ≢i j (π ⟨$⟩ʳ punchOut i≢k) ∘ sym)) ≡⟨ cong (λ l → punchIn i (π ⟨$⟩ˡ l)) (punchOut-punchIn j) ⟩
punchIn i (π ⟨$⟩ˡ (π ⟨$⟩ʳ punchOut i≢k)) ≡⟨ cong (punchIn i) (inverseˡ π) ⟩
punchIn i (punchOut i≢k) ≡⟨ punchIn-punchOut i≢k ⟩
k ∎
punchIn i (π ⟨$⟩ˡ (π ⟨$⟩ʳ punchOut i≢k)) ≡⟨ cong (punchIn i) (inverseˡ π) ⟩
punchIn i (punchOut i≢k) ≡⟨ punchIn-punchOut i≢k ⟩
k

inverseˡ′ : StrictlyInverseˡ _≡_ to from
inverseˡ′ k with j ≟ k
... | yes j≡k rewrite proj₂ (dec-yes (i ≟ i) refl) = j≡k
... | yes j≡k rewrite ≟-diag-refl i = j≡k
... | no j≢k
with i≢punchInᵢπˡpunchOutj≢k ← punchInᵢ≢i i (π ⟨$⟩ˡ punchOut j≢k) ∘ sym
rewrite dec-no (i ≟ punchIn i (π ⟨$⟩ˡ punchOut j≢k)) i≢punchInᵢπˡpunchOutj≢k
rewrite ≟-off-diag i≢punchInᵢπˡpunchOutj≢k
= begin
punchIn j (π ⟨$⟩ʳ punchOut i≢punchInᵢπˡpunchOutj≢k) ≡⟨ cong (λ l → punchIn j (π ⟨$⟩ʳ l)) (punchOut-cong i refl) ⟩
punchIn j (π ⟨$⟩ʳ punchOut (punchInᵢ≢i i (π ⟨$⟩ˡ punchOut j≢k) ∘ sym)) ≡⟨ cong (λ l → punchIn j (π ⟨$⟩ʳ l)) (punchOut-punchIn i) ⟩
Expand Down Expand Up @@ -334,17 +335,12 @@ insert-remove : ∀ i (π : Permutation (suc m) (suc n)) → insert i (π ⟨$
insert-remove {m = m} {n = n} i π j with i ≟ j
... | yes i≡j = cong (π ⟨$⟩ʳ_) i≡j
... | no i≢j = begin
punchIn (π ⟨$⟩ʳ i) (punchOut (punchInᵢ≢i i (punchOut i≢j) ∘ sym ∘ Injection.injective (↔⇒↣ π))) ≡⟨ punchIn-punchOut _ ⟩
π ⟨$⟩ʳ punchIn i (punchOut i≢j) ≡⟨ cong (π ⟨$⟩ʳ_) (punchIn-punchOut i≢j) ⟩
π ⟨$⟩ʳ j ∎
punchIn (π ⟨$⟩ʳ i) (punchOut (punchInᵢ≢i i (punchOut i≢j) ∘ sym ∘ Injection.injective (↔⇒↣ π))) ≡⟨ punchIn-punchOut _ ⟩
π ⟨$⟩ʳ punchIn i (punchOut i≢j) ≡⟨ cong (π ⟨$⟩ʳ_) (punchIn-punchOut i≢j) ⟩
π ⟨$⟩ʳ j

remove-insert : ∀ i j (π : Permutation m n) → remove i (insert i j π) ≈ π
remove-insert i j π k with i ≟ i
... | no i≢i = contradiction refl i≢i
... | yes _ = begin
punchOut {i = j} _
≡⟨ punchOut-cong j (insert-punchIn i j π k) ⟩
punchOut {i = j} (punchInᵢ≢i j (π ⟨$⟩ʳ k) ∘ sym)
≡⟨ punchOut-punchIn j ⟩
π ⟨$⟩ʳ k
remove-insert i j π k rewrite ≟-diag-refl i = begin
punchOut {i = j} _ ≡⟨ punchOut-cong j (insert-punchIn i j π k) ⟩
punchOut {i = j} (punchInᵢ≢i j (π ⟨$⟩ʳ k) ∘ sym) ≡⟨ punchOut-punchIn j ⟩
π ⟨$⟩ʳ k ∎
50 changes: 28 additions & 22 deletions src/Data/Fin/Permutation/Components.agda
Original file line number Diff line number Diff line change
Expand Up @@ -11,25 +11,17 @@ module Data.Fin.Permutation.Components where
open import Data.Bool.Base using (Bool; true; false)
open import Data.Fin.Base using (Fin; suc; opposite; toℕ)
open import Data.Fin.Properties
using (_≟_; opposite-prop; opposite-involutive; opposite-suc)
open import Data.Nat.Base as ℕ using (zero; suc; _∸_)
open import Data.Product.Base using (proj₂)
open import Function.Base using (_∘_)
open import Relation.Nullary.Reflects using (invert)
open import Relation.Nullary using (does; _because_; yes; no)
open import Relation.Nullary.Decidable using (dec-true; dec-false)
using (_≟_; ≟-diag; ≟-diag-refl
; opposite-prop; opposite-involutive; opposite-suc)
open import Relation.Nullary.Decidable.Core using (does; yes; no)
open import Relation.Binary.PropositionalEquality.Core
using (_≡_; refl; sym; trans)
open import Relation.Binary.PropositionalEquality.Properties
using (module ≡-Reasoning)
open import Algebra.Definitions using (Involutive)
open ≡-Reasoning
using (_≡_; refl; sym)

------------------------------------------------------------------------
-- Functions
------------------------------------------------------------------------

-- 'tranpose i j' swaps the places of 'i' and 'j'.
-- 'transpose i j' swaps the places of 'i' and 'j'.

transpose : ∀ {n} → Fin n → Fin n → Fin n → Fin n
transpose i j k with does (k ≟ i)
Expand All @@ -42,17 +34,31 @@ transpose i j k with does (k ≟ i)
-- Properties
------------------------------------------------------------------------

transpose-iij : ∀ {n} (i j : Fin n) → transpose i i j ≡ j
transpose-iij i j with j ≟ i in j≟i
... | yes j≡i = sym j≡i
... | no _ rewrite j≟i = refl

transpose-ijj : ∀ {n} (i j : Fin n) → transpose i j j ≡ i
transpose-ijj i j with j ≟ i
... | yes j≡i = j≡i
... | no _ rewrite ≟-diag-refl j = refl

transpose-iji : ∀ {n} (i j : Fin n) → transpose i j i ≡ j
transpose-iji i j rewrite ≟-diag-refl i = refl

transpose-transpose : ∀ {n} {i j k l : Fin n} →
transpose i j k ≡ l → transpose j i l ≡ k
transpose-transpose {n} {i} {j} {k} {l} eq with k ≟ i in k≟i
... | yes k≡i rewrite ≟-diag (sym eq) = sym k≡i
... | no k≢i with k ≟ j in k≟j
... | yes k≡j rewrite eq | transpose-ijj j l = sym k≡j
... | no k≢j rewrite eq | k≟j | k≟i = refl

transpose-inverse : ∀ {n} (i j : Fin n) {k} →
transpose i j (transpose j i k) ≡ k
transpose-inverse i j {k} with k ≟ j
... | true because [k≡j] rewrite dec-true (i ≟ i) refl = sym (invert [k≡j])
... | false because [k≢j] with k ≟ i
... | true because [k≡i]
rewrite dec-false (j ≟ i) (invert [k≢j] ∘ trans (invert [k≡i]) ∘ sym)
| dec-true (j ≟ j) refl
= sym (invert [k≡i])
... | false because [k≢i] rewrite dec-false (k ≟ i) (invert [k≢i])
| dec-false (k ≟ j) (invert [k≢j]) = refl
transpose-inverse i j = transpose-transpose refl


------------------------------------------------------------------------
-- DEPRECATED NAMES
Expand Down
17 changes: 16 additions & 1 deletion src/Data/Fin/Properties.agda
Original file line number Diff line number Diff line change
Expand Up @@ -11,6 +11,7 @@
module Data.Fin.Properties where

open import Axiom.Extensionality.Propositional
open import Axiom.UniquenessOfIdentityProofs using (module Decidable⇒UIP)
open import Algebra.Definitions using (Involutive)
open import Effect.Applicative using (RawApplicative)
open import Effect.Functor using (RawFunctor)
Expand Down Expand Up @@ -44,10 +45,12 @@ open import Relation.Binary.PropositionalEquality.Core as ≡
using (_≡_; _≢_; refl; sym; trans; cong; cong₂; subst; _≗_)
open import Relation.Binary.PropositionalEquality.Properties as ≡
using (module ≡-Reasoning)
open import Relation.Binary.PropositionalEquality as ≡
using (≡-≟-identity; ≢-≟-identity)
open import Relation.Nullary.Decidable as Dec
using (Dec; _because_; yes; no; _×-dec_; _⊎-dec_; map′)
open import Relation.Nullary.Negation.Core using (¬_; contradiction)
open import Relation.Nullary.Reflects using (Reflects; invert)
open import Relation.Nullary.Reflects using (invert)
open import Relation.Unary as U
using (U; Pred; Decidable; _⊆_; Satisfiable; Universal)
open import Relation.Unary.Properties using (U?)
Expand Down Expand Up @@ -100,6 +103,18 @@ zero ≟ suc y = no λ()
suc x ≟ zero = no λ()
suc x ≟ suc y = map′ (cong suc) suc-injective (x ≟ y)

≡-irrelevant : Irrelevant {A = Fin n} _≡_
≡-irrelevant = Decidable⇒UIP.≡-irrelevant _≟_

≟-diag : (eq : i ≡ j) → (i ≟ j) ≡ yes eq
≟-diag = ≡-≟-identity _≟_

≟-diag-refl : (i : Fin n) → (i ≟ i) ≡ yes refl
≟-diag-refl _ = ≟-diag refl

≟-off-diag : (i≢j : i ≢ j) → (i ≟ j) ≡ no i≢j
≟-off-diag = ≢-≟-identity _≟_

------------------------------------------------------------------------
-- Structures

Expand Down
6 changes: 6 additions & 0 deletions src/Data/Nat/Properties.agda
Original file line number Diff line number Diff line change
Expand Up @@ -110,6 +110,12 @@ m ≟ n = map′ (≡ᵇ⇒≡ m n) (≡⇒≡ᵇ m n) (T? (m ≡ᵇ n))
≟-diag : (eq : m ≡ n) → (m ≟ n) ≡ yes eq
≟-diag = ≡-≟-identity _≟_

≟-diag-refl : ∀ n → (n ≟ n) ≡ yes refl
≟-diag-refl _ = ≟-diag refl
Copy link
Contributor

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

This one, on the other hand, seems to fail the Fairbairn test.

Copy link
Contributor Author

@jamesmckinna jamesmckinna Jun 26, 2025

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Well, the reflexive case us what I use for Fin, and I dislike having the APIs for Nat and Fin diverging... what is really prefer would be to only have diag refer to the refl version but... legacy names!


≟-off-diag : (m≢n : m ≢ n) → (m ≟ n) ≡ no m≢n
≟-off-diag = ≢-≟-identity _≟_

≡-isDecEquivalence : IsDecEquivalence (_≡_ {A = ℕ})
≡-isDecEquivalence = record
{ isEquivalence = isEquivalence
Expand Down