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

@@ -38,8 +38,9 @@ module TransportFiniteHeight
open IsEquivalence ≈₁-equiv using () renaming (≈-sym to ≈₁-sym; ≈-trans to ≈₁-trans)
open IsEquivalence ≈₂-equiv using () renaming (≈-sym to ≈₂-sym; ≈-trans to ≈₂-trans)
open import Chain _≈₁_ ≈₁-equiv _≺₁_ ≺₁-cong using () renaming (Chain to Chain₁; done to done₁; step to step₁)
open import Chain _≈_ -equiv _≺_ -cong using () renaming (Chain to Chain; done to done; step to step)
import Chain
open Chain _≈_ -equiv _≺_ -cong using () renaming (Chain to Chain; done to done; step to step)
open Chain _≈₂_ ≈₂-equiv _≺₂_ ≺₂-cong using () renaming (Chain to Chain₂; done to done₂; step to step₂)
private
f-Injective : Injective _≈₁_ _≈₂_ f
@@ -65,10 +66,17 @@ module TransportFiniteHeight
isFiniteHeightLattice : IsFiniteHeightLattice B height _≈₂_ _⊔₂_ _⊓₂_
isFiniteHeightLattice =
let
(((a₁ , a₂) , c) , bounded₁) = IsFiniteHeightLattice.fixedHeight fhlA
open Chain.Height (IsFiniteHeightLattice.fixedHeight fhlA)
using ()
renaming ( to ⊥₁; to ⊤₁; bounded to bounded₁; longestChain to c)
in record
{ isLattice = lB
; fixedHeight = (((f a₁ , f a₂), portChain₁ c) , λ c' bounded₁ (portChain₂ c'))
; fixedHeight = record
{ = f ⊥₁
; = f ⊤₁
; longestChain = portChain₁ c
; bounded = λ c' bounded₁ (portChain₂ c')
}
}
finiteHeightLattice : FiniteHeightLattice B