open import Lattice module Lattice.Bundles.IterProd {a} (A B : Set a) where open import Data.Nat using (ℕ) module _ (lA : Lattice A) (lB : Lattice B) where open Lattice.Lattice lA using () renaming ( _≈_ to _≈₁_; _⊔_ to _⊔₁_; _⊓_ to _⊓₁_ ; isLattice to isLattice₁ ) open Lattice.Lattice lB using () renaming ( _≈_ to _≈₂_; _⊔_ to _⊔₂_; _⊓_ to _⊓₂_ ; isLattice to isLattice₂ ) module _ (k : ℕ) where open import Lattice.IterProd _≈₁_ _≈₂_ _⊔₁_ _⊔₂_ _⊓₁_ _⊓₂_ isLattice₁ isLattice₂ using (lattice) public module _ (fhA : FiniteHeightLattice A) (fhB : FiniteHeightLattice B) where open Lattice.FiniteHeightLattice fhA using () renaming ( _≈_ to _≈₁_; _⊔_ to _⊔₁_; _⊓_ to _⊓₁_ ; height to height₁ ; isLattice to isLattice₁ ; fixedHeight to fixedHeight₁ ) open Lattice.FiniteHeightLattice fhB using () renaming ( _≈_ to _≈₂_; _⊔_ to _⊔₂_; _⊓_ to _⊓₂_ ; height to height₂ ; isLattice to isLattice₂ ; fixedHeight to fixedHeight₂ ) module _ (≈₁-dec : IsDecidable _≈₁_) (≈₂-dec : IsDecidable _≈₂_) (k : ℕ) where import Lattice.IterProd _≈₁_ _≈₂_ _⊔₁_ _⊔₂_ _⊓₁_ _⊓₂_ isLattice₁ isLattice₂ as IP finiteHeightLattice = IP.finiteHeightLattice k ≈₁-dec ≈₂-dec height₁ height₂ fixedHeight₁ fixedHeight₂