diff --git a/README/Data/Fin/Relation/Ternary/Pinch.agda b/README/Data/Fin/Relation/Ternary/Pinch.agda new file mode 100644 index 0000000000..e0d2b9e453 --- /dev/null +++ b/README/Data/Fin/Relation/Ternary/Pinch.agda @@ -0,0 +1,53 @@ +------------------------------------------------------------------------ +-- The Agda standard library +-- +-- Example use of the 'Pinch' view of Fin +-- +-- This is an example of a view of a function defined over a datatype, +-- such that the recursion and call-pattern(s) of the function are +-- precisely mirrored in the constructors of the view type +-- +-- Using this view, we can exhibit the corresponding properties of +-- the function `punchIn` defined in `Data.Fin.Properties` +------------------------------------------------------------------------ + +{-# OPTIONS --cubical-compatible --safe #-} + +module README.Data.Fin.Relation.Ternary.Pinch where + +open import Data.Nat.Base as Nat using (ℕ; suc; ∣_-_∣) +open import Data.Fin.Base using (Fin; zero; suc; toℕ; _≤_; _<_; pinch) +open import Data.Fin.Relation.Ternary.Pinch +open import Data.Product using (_,_; ∃) +open import Function.Definitions.Core2 using (Surjective) +open import Relation.Binary.Core +open import Relation.Binary.PropositionalEquality using (_≡_; _≢_; refl) + +private + variable + n : ℕ + + +------------------------------------------------------------------------ +-- Properties of the function, derived from properties of the View + +pinch-surjective : ∀ (i : Fin n) → Surjective _≡_ (pinch i) +pinch-surjective i k + with j , v ← view-surjective i k + with refl ← view-complete v = j , refl + +pinch-injective : ∀ {i : Fin n} {j k : Fin (ℕ.suc n)} → + suc i ≢ j → suc i ≢ k → pinch i j ≡ pinch i k → j ≡ k +pinch-injective {n = n} {i} {j} {k} = view-injective (view i j) (view i k) + +pinch-mono-≤ : ∀ (i : Fin n) → (pinch i) Preserves _≤_ ⟶ _≤_ +pinch-mono-≤ i {j} {k} = view-mono-≤ (view i j) (view i k) + +pinch-cancel-< : ∀ (i : Fin (suc n)) {j k} → + (pinch i j < pinch i k) → j < k +pinch-cancel-< i {j} {k} = view-cancel-< (view i j) (view i k) + +pinch-≡⇒∣j-k∣≤1 : ∀ (i : Fin n) {j k} → + (pinch i j ≡ pinch i k) → ∣ (toℕ j) - (toℕ k) ∣ Nat.≤ 1 +pinch-≡⇒∣j-k∣≤1 i {j} {k} = view-j≡view-k⇒∣j-k∣≤1 (view i j) (view i k) + diff --git a/README/Data/Fin/Relation/Ternary/PunchIn.agda b/README/Data/Fin/Relation/Ternary/PunchIn.agda new file mode 100644 index 0000000000..ae04b2774b --- /dev/null +++ b/README/Data/Fin/Relation/Ternary/PunchIn.agda @@ -0,0 +1,44 @@ +------------------------------------------------------------------------ +-- The Agda standard library +-- +-- Example use of the 'PunchIn' view of Fin +-- +-- This is an example of a view of a function defined over a datatype, +-- such that the recursion and call-pattern(s) of the function are +-- precisely mirrored in the constructors of the view type +-- +-- Using this view, we can exhibit the corresponding properties of +-- the function `punchIn` defined in `Data.Fin.Properties` +------------------------------------------------------------------------ + +{-# OPTIONS --cubical-compatible --safe #-} + +module README.Data.Fin.Relation.Ternary.PunchIn where + +open import Data.Nat.Base using (ℕ; suc) +open import Data.Fin.Base using (Fin; zero; suc; _≤_; punchIn) +open import Data.Fin.Relation.Ternary.PunchIn +open import Function.Definitions using (Injective) +open import Relation.Binary.Core +open import Relation.Binary.PropositionalEquality using (_≡_; _≢_) + +private + variable + n : ℕ + +------------------------------------------------------------------------ +-- Properties of the function, derived from properties of the View + +punchInᵢ≢i : ∀ i (j : Fin n) → punchIn i j ≢ i +punchInᵢ≢i i j = view-codomain (view i j) + +punchIn-injective : ∀ (i : Fin (suc n)) → Injective _≡_ _≡_ (punchIn i) +punchIn-injective i {j} {k} = view-injective (view i j) (view i k) + +punchIn-mono-≤ : ∀ (i : Fin (suc n)) → (punchIn i) Preserves _≤_ ⟶ _≤_ +punchIn-mono-≤ i {j} {k} = view-mono-≤ (view i j) (view i k) + +punchIn-cancel-≤ : ∀ (i : Fin (suc n)) {j k} → + (punchIn i j ≤ punchIn i k) → j ≤ k +punchIn-cancel-≤ i {j} {k} = view-cancel-≤ (view i j) (view i k) + diff --git a/README/Data/Fin/Relation/Ternary/PunchOut.agda b/README/Data/Fin/Relation/Ternary/PunchOut.agda new file mode 100644 index 0000000000..557c7ddc7c --- /dev/null +++ b/README/Data/Fin/Relation/Ternary/PunchOut.agda @@ -0,0 +1,59 @@ +------------------------------------------------------------------------ +-- The Agda standard library +-- +-- Example use of the 'PunchOut' view of Fin +-- +-- This is an example of a view of a function defined over a datatype, +-- such that the recursion and call-pattern(s) of the function are +-- precisely mirrored in the constructors of the view type +-- +-- Using this view, we can exhibit the corresponding properties of +-- the function `punchOut` defined in `Data.Fin.Properties` +------------------------------------------------------------------------ + +{-# OPTIONS --cubical-compatible --safe #-} + +module README.Data.Fin.Relation.Ternary.PunchOut where + +open import Data.Nat.Base using (ℕ; suc) +open import Data.Fin.Base using (Fin; zero; suc; _≤_; punchIn; punchOut) +open import Data.Fin.Properties using (punchInᵢ≢i) +open import Data.Fin.Relation.Ternary.PunchIn as PunchIn using () +open import Data.Fin.Relation.Ternary.PunchOut +open import Function.Base using (_∘_) +open import Relation.Binary.Core +open import Relation.Binary.PropositionalEquality using (_≡_; _≢_; sym) + +private + variable + n : ℕ + +------------------------------------------------------------------------ +-- Properties of the function, derived from properties of the View + +punchOut-injective : ∀ {i j k : Fin (suc n)} (i≢j : i ≢ j) (i≢k : i ≢ k) → + punchOut i≢j ≡ punchOut i≢k → j ≡ k +punchOut-injective i≢j i≢k = view-injective (view i≢j) (view i≢k) + +punchOut-cong : ∀ (i : Fin (suc n)) {j k} {i≢j : i ≢ j} {i≢k : i ≢ k} → + j ≡ k → punchOut i≢j ≡ punchOut i≢k +punchOut-cong i {i≢j = i≢j} {i≢k = i≢k} = view-cong (view i≢j) (view i≢k) + +punchOut-mono-≤ : ∀ {i j k : Fin (suc n)} (i≢j : i ≢ j) (i≢k : i ≢ k) → + j ≤ k → punchOut i≢j ≤ punchOut i≢k +punchOut-mono-≤ i≢j i≢k = view-mono-≤ (view i≢j) (view i≢k) + +punchOut-cancel-≤ : ∀ {i j k : Fin (suc n)} (i≢j : i ≢ j) (i≢k : i ≢ k) → + (punchOut i≢j ≤ punchOut i≢k) → j ≤ k +punchOut-cancel-≤ i≢j i≢k = view-cancel-≤ (view i≢j) (view i≢k) + +-- punchOut and punchIn are mutual inverses, +-- because their corresponding View s are converses + +punchIn-punchOut : ∀ {i j : Fin (suc n)} (i≢j : i ≢ j) → + punchIn i (punchOut i≢j) ≡ j +punchIn-punchOut = PunchIn.view-complete ∘ view-view⁻¹ ∘ view + +punchOut-punchIn : ∀ i {j : Fin n} → + punchOut {i = i} {j = punchIn i j} (punchInᵢ≢i i j ∘ sym) ≡ j +punchOut-punchIn i {j} = view-complete (view⁻¹-view (PunchIn.view i j)) diff --git a/src/Data/Fin/Relation/Ternary/Pinch.agda b/src/Data/Fin/Relation/Ternary/Pinch.agda new file mode 100644 index 0000000000..26db16461c --- /dev/null +++ b/src/Data/Fin/Relation/Ternary/Pinch.agda @@ -0,0 +1,119 @@ +------------------------------------------------------------------------ +-- The Agda standard library +-- +-- The '`Pinch` view' of the function `pinch` defined on finite sets +------------------------------------------------------------------------ +-- +-- This is an example of a view of a function defined over a datatype, +-- such that the recursion and call-pattern(s) of the function are +-- precisely mirrored in the constructors of the view type, +-- ie that we 'view the function via its graph relation' + +{-# OPTIONS --cubical-compatible --safe #-} + +module Data.Fin.Relation.Ternary.Pinch where + +open import Data.Fin.Base using (Fin; zero; suc; toℕ; _≤_; _<_; pinch) +open import Data.Fin.Properties using (suc-injective) +open import Data.Nat.Base as Nat using (ℕ; zero; suc; z≤n; s≤s; zi then j-1 else j + +punchOut : ∀ {i j : Fin (suc n)} → i ≢ j → Fin n +punchOut {_} {zero} {zero} i≢j = ⊥-elim (i≢j refl) +punchOut {_} {zero} {suc j} _ = j +punchOut {suc _} {suc i} {zero} _ = zero +punchOut {suc _} {suc i} {suc j} i≢j = suc (punchOut (i≢j ∘ cong suc)) + +-} + zero-suc : ∀ {n} (j : Fin n) → View zero (suc j) j + suc-zero : ∀ {n} (i : Fin (suc n)) → View (suc i) zero zero + suc-suc : ∀ {n} {i} {j} {k} → View {n} i j k → View (suc i) (suc j) (suc k) + +-- The View enforces the precondition given by a Domain predicate + +Domain : ∀ (i j : Fin (suc n)) → Set +Domain i j = i ≢ j + +view-domain : ∀ {i j} {k} → View {n} i j k → Domain i j +view-domain (suc-suc v) = (view-domain v) ∘ suc-injective + +-- The View is sound, ie covers all telescopes satisfying that precondition + +-- The recursion/pattern analysis of the original definition of `punchOut` +-- (which is responsible for showing termination in the first place) +-- is then exactly replicated in the definition of the covering function `view`; +-- thus that definitional pattern analysis is encapsulated once and for all + +view : ∀ {n} {i j} (d : Domain i j) → View {n} i j (punchOut d) +view {i = zero} {j = zero} d with () ← d refl +view {i = zero} {j = suc j} d = zero-suc j +view {n = suc _} {i = suc i} {j = zero} d = suc-zero i +view {n = suc _} {i = suc i} {j = suc j} d = suc-suc (view (d ∘ (cong suc))) + +-- Interpretation of the view: the original function itself + +⟦_⟧ : ∀ {i j} {k} (v : View {n} i j k) → Fin n +⟦ v ⟧ = punchOut (view-domain v) + +-- The View is complete + +view-complete : ∀ {i j} {k} (v : View {n} i j k) → ⟦ v ⟧ ≡ k +view-complete (zero-suc j) = refl +view-complete (suc-zero i) = refl +view-complete (suc-suc v) = cong suc (view-complete v) + +------------------------------------------------------------------------ +-- Properties of the function, derived from properties of the View + +view-cong : ∀ {i j k} {p q} → View {n} i j p → View {n} i k q → + j ≡ k → p ≡ q +view-cong v w refl = aux v w where + aux : ∀ {i j} {p q} → View {n} i j p → View {n} i j q → p ≡ q + aux (zero-suc _) (zero-suc _) = refl + aux (suc-zero i) (suc-zero i) = refl + aux (suc-suc v) (suc-suc w) = cong suc (aux v w) + +view-injective : ∀ {i j k} {p q} → View {n} i j p → View {n} i k q → + p ≡ q → j ≡ k +view-injective v w refl = aux v w where + aux : ∀ {i j k} {r} → View {n} i j r → View {n} i k r → j ≡ k + aux (zero-suc _) (zero-suc _) = refl + aux (suc-zero i) (suc-zero i) = refl + aux (suc-suc v) (suc-suc w) = cong suc (aux v w) + +view-mono-≤ : ∀ {i j k} {p q} → View {n} i j p → View {n} i k q → + j ≤ k → p ≤ q +view-mono-≤ (zero-suc j) (zero-suc k) (s≤s j≤k) = j≤k +view-mono-≤ (suc-zero i) _ _ = z≤n +view-mono-≤ (suc-suc vj) (suc-suc vk) (s≤s j≤k) = s≤s (view-mono-≤ vj vk j≤k) + +view-cancel-≤ : ∀ {i j k} {p q} → View {n} i j p → View {n} i k q → + p ≤ q → j ≤ k +view-cancel-≤ (zero-suc j) (zero-suc k) p≤q = s≤s p≤q +view-cancel-≤ (suc-zero i) _ _ = z≤n +view-cancel-≤ (suc-suc vj) (suc-suc vk) (s≤s p≤q) = s≤s (view-cancel-≤ vj vk p≤q) + +-- PunchOut.View and PunchIn.View are mutual converses + +view-view⁻¹ : ∀ {i j} {k} → View {n} i j k → PunchIn.View {n} i k j +view-view⁻¹ (zero-suc _) = PunchIn.zero-suc _ +view-view⁻¹ (suc-zero i) = PunchIn.suc-zero i +view-view⁻¹ (suc-suc v) = PunchIn.suc-suc (view-view⁻¹ v) + +view⁻¹-view : ∀ {i j} {k} → PunchIn.View {n} i k j → View {n} i j k +view⁻¹-view (PunchIn.zero-suc _) = zero-suc _ +view⁻¹-view (PunchIn.suc-zero i) = suc-zero i +view⁻¹-view (PunchIn.suc-suc v) = suc-suc (view⁻¹-view v)