2024-02-11 14:18:58 -08:00
|
|
|
|
open import Lattice
|
|
|
|
|
|
2024-02-19 12:36:46 -08:00
|
|
|
|
-- Due to universe levels, it becomes relatively annoying to handle the case
|
|
|
|
|
-- where the levels of A and B are not the same. For the time being, constrain
|
|
|
|
|
-- them to be the same.
|
|
|
|
|
|
2024-02-11 14:18:58 -08:00
|
|
|
|
module Lattice.IterProd {a} {A B : Set a}
|
|
|
|
|
(_≈₁_ : A → A → Set a) (_≈₂_ : B → B → Set a)
|
|
|
|
|
(_⊔₁_ : A → A → A) (_⊔₂_ : B → B → B)
|
|
|
|
|
(_⊓₁_ : A → A → A) (_⊓₂_ : B → B → B)
|
|
|
|
|
(lA : IsLattice A _≈₁_ _⊔₁_ _⊓₁_) (lB : IsLattice B _≈₂_ _⊔₂_ _⊓₂_) where
|
|
|
|
|
|
2024-02-11 15:36:12 -08:00
|
|
|
|
open import Agda.Primitive using (lsuc)
|
|
|
|
|
open import Data.Nat using (ℕ; suc; _+_)
|
2024-02-11 14:18:58 -08:00
|
|
|
|
open import Data.Product using (_×_)
|
|
|
|
|
open import Utils using (iterate)
|
|
|
|
|
|
2024-02-11 15:36:12 -08:00
|
|
|
|
open IsLattice lA renaming (FixedHeight to FixedHeight₁)
|
|
|
|
|
open IsLattice lB renaming (FixedHeight to FixedHeight₂)
|
|
|
|
|
|
2024-02-11 14:18:58 -08:00
|
|
|
|
IterProd : ℕ → Set a
|
|
|
|
|
IterProd k = iterate k (λ t → A × t) B
|
|
|
|
|
|
|
|
|
|
-- To make iteration more convenient, package the definitions in Lattice
|
|
|
|
|
-- records, perform the recursion, and unpackage.
|
|
|
|
|
|
2024-02-18 20:18:38 -08:00
|
|
|
|
module _ where
|
2024-02-11 21:10:51 -08:00
|
|
|
|
lattice : ∀ {k : ℕ} → Lattice (IterProd k)
|
|
|
|
|
lattice {0} = record
|
2024-02-11 14:18:58 -08:00
|
|
|
|
{ _≈_ = _≈₂_
|
|
|
|
|
; _⊔_ = _⊔₂_
|
|
|
|
|
; _⊓_ = _⊓₂_
|
|
|
|
|
; isLattice = lB
|
|
|
|
|
}
|
2024-02-11 21:10:51 -08:00
|
|
|
|
lattice {suc k'} = record
|
2024-02-11 14:18:58 -08:00
|
|
|
|
{ _≈_ = _≈_
|
|
|
|
|
; _⊔_ = _⊔_
|
|
|
|
|
; _⊓_ = _⊓_
|
|
|
|
|
; isLattice = isLattice
|
|
|
|
|
}
|
|
|
|
|
where
|
2024-02-11 15:36:12 -08:00
|
|
|
|
Right : Lattice (IterProd k')
|
2024-02-11 21:10:51 -08:00
|
|
|
|
Right = lattice {k'}
|
2024-02-11 14:18:58 -08:00
|
|
|
|
|
|
|
|
|
open import Lattice.Prod
|
2024-02-11 15:36:12 -08:00
|
|
|
|
_≈₁_ (Lattice._≈_ Right)
|
|
|
|
|
_⊔₁_ (Lattice._⊔_ Right)
|
|
|
|
|
_⊓₁_ (Lattice._⊓_ Right)
|
|
|
|
|
lA (Lattice.isLattice Right)
|
|
|
|
|
|
2024-02-11 20:45:14 -08:00
|
|
|
|
module _ (k : ℕ) where
|
2024-02-11 21:10:51 -08:00
|
|
|
|
open Lattice.Lattice (lattice {k}) public
|
2024-02-11 20:45:14 -08:00
|
|
|
|
|
2024-02-11 15:36:12 -08:00
|
|
|
|
module _ (≈₁-dec : IsDecidable _≈₁_) (≈₂-dec : IsDecidable _≈₂_)
|
|
|
|
|
(h₁ h₂ : ℕ)
|
|
|
|
|
(fhA : FixedHeight₁ h₁) (fhB : FixedHeight₂ h₂) where
|
|
|
|
|
|
|
|
|
|
private module _ where
|
|
|
|
|
record FiniteHeightAndDecEq (A : Set a) : Set (lsuc a) where
|
|
|
|
|
field
|
|
|
|
|
height : ℕ
|
|
|
|
|
_≈_ : A → A → Set a
|
|
|
|
|
_⊔_ : A → A → A
|
|
|
|
|
_⊓_ : A → A → A
|
|
|
|
|
|
|
|
|
|
isFiniteHeightLattice : IsFiniteHeightLattice A height _≈_ _⊔_ _⊓_
|
|
|
|
|
≈-dec : IsDecidable _≈_
|
|
|
|
|
|
|
|
|
|
open IsFiniteHeightLattice isFiniteHeightLattice public
|
|
|
|
|
|
2024-02-11 21:10:51 -08:00
|
|
|
|
finiteHeightAndDec : ∀ {k : ℕ} → FiniteHeightAndDecEq (IterProd k)
|
|
|
|
|
finiteHeightAndDec {0} = record
|
2024-02-11 15:36:12 -08:00
|
|
|
|
{ height = h₂
|
|
|
|
|
; _≈_ = _≈₂_
|
|
|
|
|
; _⊔_ = _⊔₂_
|
|
|
|
|
; _⊓_ = _⊓₂_
|
|
|
|
|
; isFiniteHeightLattice = record
|
|
|
|
|
{ isLattice = lB
|
|
|
|
|
; fixedHeight = fhB
|
|
|
|
|
}
|
|
|
|
|
; ≈-dec = ≈₂-dec
|
|
|
|
|
}
|
2024-02-11 21:10:51 -08:00
|
|
|
|
finiteHeightAndDec {suc k'} = record
|
2024-02-11 15:36:12 -08:00
|
|
|
|
{ height = h₁ + FiniteHeightAndDecEq.height Right
|
2024-02-11 20:45:14 -08:00
|
|
|
|
; _≈_ = P._≈_
|
|
|
|
|
; _⊔_ = P._⊔_
|
|
|
|
|
; _⊓_ = P._⊓_
|
2024-02-11 15:36:12 -08:00
|
|
|
|
; isFiniteHeightLattice = isFiniteHeightLattice
|
|
|
|
|
≈₁-dec (FiniteHeightAndDecEq.≈-dec Right)
|
|
|
|
|
h₁ (FiniteHeightAndDecEq.height Right)
|
|
|
|
|
fhA (IsFiniteHeightLattice.fixedHeight (FiniteHeightAndDecEq.isFiniteHeightLattice Right))
|
|
|
|
|
; ≈-dec = ≈-dec ≈₁-dec (FiniteHeightAndDecEq.≈-dec Right)
|
|
|
|
|
}
|
|
|
|
|
where
|
2024-02-11 21:10:51 -08:00
|
|
|
|
Right = finiteHeightAndDec {k'}
|
2024-02-11 15:36:12 -08:00
|
|
|
|
|
|
|
|
|
open import Lattice.Prod
|
|
|
|
|
_≈₁_ (FiniteHeightAndDecEq._≈_ Right)
|
|
|
|
|
_⊔₁_ (FiniteHeightAndDecEq._⊔_ Right)
|
|
|
|
|
_⊓₁_ (FiniteHeightAndDecEq._⊓_ Right)
|
2024-02-11 20:45:14 -08:00
|
|
|
|
lA (FiniteHeightAndDecEq.isLattice Right) as P
|
2024-02-11 15:36:12 -08:00
|
|
|
|
|
|
|
|
|
module _ (k : ℕ) where
|
2024-02-11 21:10:51 -08:00
|
|
|
|
open FiniteHeightAndDecEq (finiteHeightAndDec {k}) using (isFiniteHeightLattice) public
|
2024-02-11 14:18:58 -08:00
|
|
|
|
|
2024-02-11 20:45:14 -08:00
|
|
|
|
private
|
2024-02-11 21:10:51 -08:00
|
|
|
|
FHD = finiteHeightAndDec {k}
|
2024-02-11 20:45:14 -08:00
|
|
|
|
|
|
|
|
|
finiteHeightLattice : FiniteHeightLattice (IterProd k)
|
|
|
|
|
finiteHeightLattice = record
|
|
|
|
|
{ height = FiniteHeightAndDecEq.height FHD
|
|
|
|
|
; _≈_ = FiniteHeightAndDecEq._≈_ FHD
|
|
|
|
|
; _⊔_ = FiniteHeightAndDecEq._⊔_ FHD
|
|
|
|
|
; _⊓_ = FiniteHeightAndDecEq._⊓_ FHD
|
|
|
|
|
; isFiniteHeightLattice = isFiniteHeightLattice
|
|
|
|
|
}
|