Use records rather than nested pairs to represent 'fixed height'

Signed-off-by: Danila Fedorin <danila.fedorin@gmail.com>
This commit is contained in:
2024-05-09 20:11:04 -07:00
parent 95669b2c65
commit 16fa4cd1d8
8 changed files with 62 additions and 39 deletions

View File

@@ -10,7 +10,7 @@ open import Data.Nat using (; suc; _+_; _≤_)
open import Data.Nat.Properties using (+-comm; m+1+n≰m)
open import Data.Product using (_×_; Σ; _,_)
open import Relation.Binary.PropositionalEquality as Eq using (_≡_; refl)
open import Data.Empty using ()
open import Data.Empty as Empty using ()
open IsEquivalence ≈-equiv
@@ -38,11 +38,16 @@ module _ where
Bounded : Set a
Bounded bound = {a₁ a₂ : A} {n : } Chain a₁ a₂ n n bound
Bounded-suc-n : {a₁ a₂ : A} {n : } Bounded n Chain a₁ a₂ (suc n)
Bounded-suc-n : {a₁ a₂ : A} {n : } Bounded n Chain a₁ a₂ (suc n) Empty.
Bounded-suc-n {a₁} {a₂} {n} bounded c = (m+1+n≰m n n+1≤n)
where
n+1≤n : n + 1 n
n+1≤n rewrite (+-comm n 1) = bounded c
Height : Set a
Height height = (Σ (A × A) (λ (a₁ , a₂) Chain a₁ a₂ height) × Bounded height)
record Height (height : ) : Set a where
field
: A
: A
longestChain : Chain height
bounded : Bounded height