open import Lattice -- 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. 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 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₂) 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. -- -- If we prove isLattice and isFiniteHeightLattice by induction separately, -- we lose the connection between the operations (which should be the same) -- that are built up by the two iterations. So, do everything in one iteration. -- This requires some odd code. build : A → B → (k : ℕ) → IterProd k build _ b zero = b build a b (suc s) = (a , build a b s) private record RequiredForFixedHeight : Set (lsuc a) where field ≈₁-dec : IsDecidable _≈₁_ ≈₂-dec : IsDecidable _≈₂_ h₁ h₂ : ℕ fhA : FixedHeight₁ h₁ fhB : FixedHeight₂ h₂ ⊥₁ : A ⊥₁ = Height.⊥ fhA ⊥₂ : B ⊥₂ = Height.⊥ fhB ⊥k : ∀ (k : ℕ) → IterProd k ⊥k = build ⊥₁ ⊥₂ record IsFiniteHeightWithBotAndDecEq {A : Set a} {_≈_ : A → A → Set a} {_⊔_ : A → A → A} {_⊓_ : A → A → A} (isLattice : IsLattice A _≈_ _⊔_ _⊓_) (⊥ : A) : Set (lsuc a) where field height : ℕ fixedHeight : IsLattice.FixedHeight isLattice height ≈-dec : IsDecidable _≈_ ⊥-correct : Height.⊥ fixedHeight ≈ ⊥ record Everything (k : ℕ) : Set (lsuc a) where T = IterProd k field _≈_ : T → T → Set a _⊔_ : T → T → T _⊓_ : T → T → T isLattice : IsLattice T _≈_ _⊔_ _⊓_ isFiniteHeightIfSupported : ∀ (req : RequiredForFixedHeight) → IsFiniteHeightWithBotAndDecEq isLattice (RequiredForFixedHeight.⊥k req k) everything : ∀ (k : ℕ) → Everything k everything 0 = record { _≈_ = _≈₂_ ; _⊔_ = _⊔₂_ ; _⊓_ = _⊓₂_ ; isLattice = lB ; isFiniteHeightIfSupported = λ req → record { height = RequiredForFixedHeight.h₂ req ; fixedHeight = RequiredForFixedHeight.fhB req ; ≈-dec = RequiredForFixedHeight.≈₂-dec req ; ⊥-correct = IsLattice.≈-refl lB } } everything (suc k') = record { _≈_ = P._≈_ ; _⊔_ = P._⊔_ ; _⊓_ = P._⊓_ ; isLattice = P.isLattice ; isFiniteHeightIfSupported = λ req → let fhlRest = Everything.isFiniteHeightIfSupported everythingRest req in record { height = (RequiredForFixedHeight.h₁ req) + IsFiniteHeightWithBotAndDecEq.height fhlRest ; fixedHeight = P.fixedHeight (RequiredForFixedHeight.≈₁-dec req) (IsFiniteHeightWithBotAndDecEq.≈-dec fhlRest) (RequiredForFixedHeight.h₁ req) (IsFiniteHeightWithBotAndDecEq.height fhlRest) (RequiredForFixedHeight.fhA req) (IsFiniteHeightWithBotAndDecEq.fixedHeight fhlRest) ; ≈-dec = P.≈-dec (RequiredForFixedHeight.≈₁-dec req) (IsFiniteHeightWithBotAndDecEq.≈-dec fhlRest) ; ⊥-correct = (IsLattice.≈-refl lA , IsFiniteHeightWithBotAndDecEq.⊥-correct fhlRest) } } where everythingRest = everything k' import Lattice.Prod _≈₁_ (Everything._≈_ everythingRest) _⊔₁_ (Everything._⊔_ everythingRest) _⊓₁_ (Everything._⊓_ everythingRest) lA (Everything.isLattice everythingRest) as P module _ (k : ℕ) where open Everything (everything k) using (_≈_; _⊔_; _⊓_; isLattice) public open Lattice.IsLattice isLattice public lattice : Lattice (IterProd k) lattice = record { _≈_ = _≈_ ; _⊔_ = _⊔_ ; _⊓_ = _⊓_ ; isLattice = isLattice } module _ (≈₁-dec : IsDecidable _≈₁_) (≈₂-dec : IsDecidable _≈₂_) (h₁ h₂ : ℕ) (fhA : FixedHeight₁ h₁) (fhB : FixedHeight₂ h₂) where private required : RequiredForFixedHeight required = record { ≈₁-dec = ≈₁-dec ; ≈₂-dec = ≈₂-dec ; h₁ = h₁ ; h₂ = h₂ ; fhA = fhA ; fhB = fhB } fixedHeight = IsFiniteHeightWithBotAndDecEq.fixedHeight (Everything.isFiniteHeightIfSupported (everything k) required) isFiniteHeightLattice = record { isLattice = isLattice ; fixedHeight = fixedHeight } finiteHeightLattice : FiniteHeightLattice (IterProd k) finiteHeightLattice = record { height = IsFiniteHeightWithBotAndDecEq.height (Everything.isFiniteHeightIfSupported (everything k) required) ; _≈_ = _≈_ ; _⊔_ = _⊔_ ; _⊓_ = _⊓_ ; isFiniteHeightLattice = isFiniteHeightLattice }