Use records rather than nested pairs to represent 'fixed height'
Signed-off-by: Danila Fedorin <danila.fedorin@gmail.com>
This commit is contained in:
@@ -355,7 +355,12 @@ module Plain (x : A) where
|
||||
rewrite [x]≺y⇒y≡⊤ _ _ [x]≺y with ≈-⊤-⊤ ← y≈z = ⊥-elim (¬-Chain-⊤ c)
|
||||
|
||||
fixedHeight : IsLattice.FixedHeight isLattice 2
|
||||
fixedHeight = (((⊥ , ⊤) , longestChain) , isLongest)
|
||||
fixedHeight = record
|
||||
{ ⊥ = ⊥
|
||||
; ⊤ = ⊤
|
||||
; longestChain = longestChain
|
||||
; bounded = isLongest
|
||||
}
|
||||
|
||||
isFiniteHeightLattice : IsFiniteHeightLattice AboveBelow 2 _≈_ _⊔_ _⊓_
|
||||
isFiniteHeightLattice = record
|
||||
|
||||
@@ -14,6 +14,7 @@ open import Agda.Primitive using (lsuc)
|
||||
open import Data.Nat using (ℕ; zero; suc; _+_)
|
||||
open import Data.Product using (_×_; _,_; proj₁; proj₂)
|
||||
open import Utils using (iterate)
|
||||
open import Chain using (Height)
|
||||
|
||||
open IsLattice lA renaming (FixedHeight to FixedHeight₁)
|
||||
open IsLattice lB renaming (FixedHeight to FixedHeight₂)
|
||||
@@ -44,10 +45,10 @@ private
|
||||
fhB : FixedHeight₂ h₂
|
||||
|
||||
⊥₁ : A
|
||||
⊥₁ = proj₁ (proj₁ (proj₁ fhA))
|
||||
⊥₁ = Height.⊥ fhA
|
||||
|
||||
⊥₂ : B
|
||||
⊥₂ = proj₁ (proj₁ (proj₁ fhB))
|
||||
⊥₂ = Height.⊥ fhB
|
||||
|
||||
⊥k : ∀ (k : ℕ) → IterProd k
|
||||
⊥k = build ⊥₁ ⊥₂
|
||||
@@ -58,7 +59,7 @@ private
|
||||
fixedHeight : IsLattice.FixedHeight isLattice height
|
||||
≈-dec : IsDecidable _≈_
|
||||
|
||||
⊥-correct : proj₁ (proj₁ (proj₁ fixedHeight)) ≈ ⊥
|
||||
⊥-correct : Height.⊥ fixedHeight ≈ ⊥
|
||||
|
||||
record Everything (k : ℕ) : Set (lsuc a) where
|
||||
T = IterProd k
|
||||
|
||||
@@ -143,17 +143,8 @@ module _ (≈₁-dec : IsDecidable _≈₁_) (≈₂-dec : IsDecidable _≈₂_)
|
||||
∙,b-Preserves-≈₁ : ∀ (b : B) → (λ a → (a , b)) Preserves _≈₁_ ⟶ _≈_
|
||||
∙,b-Preserves-≈₁ b {a₁} {a₂} a₁≈a₂ = (a₁≈a₂ , ≈₂-refl)
|
||||
|
||||
⊥₁ : A
|
||||
⊥₁ = proj₁ (proj₁ (proj₁ fhA))
|
||||
|
||||
⊤₁ : A
|
||||
⊤₁ = proj₂ (proj₁ (proj₁ fhA))
|
||||
|
||||
⊥₂ : B
|
||||
⊥₂ = proj₁ (proj₁ (proj₁ fhB))
|
||||
|
||||
⊤₂ : B
|
||||
⊤₂ = proj₂ (proj₁ (proj₁ fhB))
|
||||
open ChainA.Height fhA using () renaming (⊥ to ⊥₁; ⊤ to ⊤₁; longestChain to longestChain₁; bounded to bounded₁)
|
||||
open ChainB.Height fhB using () renaming (⊥ to ⊥₂; ⊤ to ⊤₂; longestChain to longestChain₂; bounded to bounded₂)
|
||||
|
||||
unzip : ∀ {a₁ a₂ : A} {b₁ b₂ : B} {n : ℕ} → Chain (a₁ , b₁) (a₂ , b₂) n → Σ (ℕ × ℕ) (λ (n₁ , n₂) → ((Chain₁ a₁ a₂ n₁ × Chain₂ b₁ b₂ n₂) × (n ≤ n₁ + n₂)))
|
||||
unzip (done (a₁≈a₂ , b₁≈b₂)) = ((0 , 0) , ((done₁ a₁≈a₂ , done₂ b₁≈b₂) , ≤-refl))
|
||||
@@ -172,15 +163,16 @@ module _ (≈₁-dec : IsDecidable _≈₁_) (≈₂-dec : IsDecidable _≈₂_)
|
||||
))
|
||||
|
||||
fixedHeight : IsLattice.FixedHeight isLattice (h₁ + h₂)
|
||||
fixedHeight =
|
||||
( ( ((⊥₁ , ⊥₂) , (⊤₁ , ⊤₂))
|
||||
, concat
|
||||
(ChainMapping₁.Chain-map (λ a → (a , ⊥₂)) (∙,b-Monotonic _) proj₁ (∙,b-Preserves-≈₁ _) (proj₂ (proj₁ fhA)))
|
||||
(ChainMapping₂.Chain-map (λ b → (⊤₁ , b)) (a,∙-Monotonic _) proj₂ (a,∙-Preserves-≈₂ _) (proj₂ (proj₁ fhB)))
|
||||
)
|
||||
, λ a₁b₁a₂b₂ → let ((n₁ , n₂) , ((a₁a₂ , b₁b₂) , n≤n₁+n₂)) = unzip a₁b₁a₂b₂
|
||||
in ≤-trans n≤n₁+n₂ (+-mono-≤ (proj₂ fhA a₁a₂) (proj₂ fhB b₁b₂))
|
||||
)
|
||||
fixedHeight = record
|
||||
{ ⊥ = (⊥₁ , ⊥₂)
|
||||
; ⊤ = (⊤₁ , ⊤₂)
|
||||
; longestChain = concat
|
||||
(ChainMapping₁.Chain-map (λ a → (a , ⊥₂)) (∙,b-Monotonic _) proj₁ (∙,b-Preserves-≈₁ _) longestChain₁)
|
||||
(ChainMapping₂.Chain-map (λ b → (⊤₁ , b)) (a,∙-Monotonic _) proj₂ (a,∙-Preserves-≈₂ _) longestChain₂)
|
||||
; bounded = λ a₁b₁a₂b₂ →
|
||||
let ((n₁ , n₂) , ((a₁a₂ , b₁b₂) , n≤n₁+n₂)) = unzip a₁b₁a₂b₂
|
||||
in ≤-trans n≤n₁+n₂ (+-mono-≤ (bounded₁ a₁a₂) (bounded₂ b₁b₂))
|
||||
}
|
||||
|
||||
isFiniteHeightLattice : IsFiniteHeightLattice (A × B) (h₁ + h₂) _≈_ _⊔_ _⊓_
|
||||
isFiniteHeightLattice = record
|
||||
|
||||
@@ -108,7 +108,12 @@ private
|
||||
isLongest (done _) = z≤n
|
||||
|
||||
fixedHeight : IsLattice.FixedHeight isLattice 0
|
||||
fixedHeight = (((tt , tt) , longestChain) , isLongest)
|
||||
fixedHeight = record
|
||||
{ ⊥ = tt
|
||||
; ⊤ = tt
|
||||
; longestChain = longestChain
|
||||
; bounded = isLongest
|
||||
}
|
||||
|
||||
isFiniteHeightLattice : IsFiniteHeightLattice ⊤ 0 _≈_ _⊔_ _⊓_
|
||||
isFiniteHeightLattice = record
|
||||
|
||||
Reference in New Issue
Block a user