agda-spa/Lattice/Bundles/IterProd.agda
2024-03-01 21:58:58 -08:00

39 lines
1.5 KiB
Agda
Raw Blame History

This file contains ambiguous Unicode characters

This file contains Unicode characters that might be confused with other characters. If you think that this is intentional, you can safely ignore this warning. Use the Escape button to reveal them.

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₂