{-# OPTIONS --cubical-compatible --safe #-}
open import Algebra.Core
open import Data.Product.Base using (proj₁; proj₂)
open import Level using (_⊔_)
open import Relation.Binary.Core using (Rel)
open import Relation.Binary.Bundles using (Setoid)
open import Relation.Binary.Structures using (IsEquivalence)
module Algebra.Lattice.Structures
  {a ℓ} {A : Set a}  
  (_≈_ : Rel A ℓ)    
  where
open import Algebra.Definitions _≈_
open import Algebra.Structures _≈_
record IsSemilattice (∙ : Op₂ A) : Set (a ⊔ ℓ) where
  field
    isBand : IsBand ∙
    comm   : Commutative ∙
  open IsBand isBand public
IsMeetSemilattice = IsSemilattice
module IsMeetSemilattice {∧} (L : IsMeetSemilattice ∧) where
  open IsSemilattice L public
    renaming
    ( ∙-cong  to ∧-cong
    ; ∙-congˡ to ∧-congˡ
    ; ∙-congʳ to ∧-congʳ
    )
IsJoinSemilattice = IsSemilattice
module IsJoinSemilattice {∨} (L : IsJoinSemilattice ∨) where
  open IsSemilattice L public
    renaming
    ( ∙-cong  to ∨-cong
    ; ∙-congˡ to ∨-congˡ
    ; ∙-congʳ to ∨-congʳ
    )
IsBoundedSemilattice = IsIdempotentCommutativeMonoid
module IsBoundedSemilattice {∙ ε} (L : IsBoundedSemilattice ∙ ε) where
  open IsIdempotentCommutativeMonoid L public
  isSemilattice : IsSemilattice ∙
  isSemilattice = record
    { isBand = isBand
    ; comm   = comm
    }
IsBoundedMeetSemilattice = IsBoundedSemilattice
module IsBoundedMeetSemilattice {∧ ⊤} (L : IsBoundedMeetSemilattice ∧ ⊤)
  where
  open IsBoundedSemilattice L public
    using (identity; identityˡ; identityʳ)
    renaming (isSemilattice to isMeetSemilattice)
  open IsMeetSemilattice isMeetSemilattice public
IsBoundedJoinSemilattice = IsBoundedSemilattice
module IsBoundedJoinSemilattice {∨ ⊥} (L : IsBoundedJoinSemilattice ∨ ⊥)
  where
  open IsBoundedSemilattice L public
    using (identity; identityˡ; identityʳ)
    renaming (isSemilattice to isJoinSemilattice)
  open IsJoinSemilattice isJoinSemilattice public
record IsLattice (∨ ∧ : Op₂ A) : Set (a ⊔ ℓ) where
  field
    isEquivalence : IsEquivalence _≈_
    ∨-comm        : Commutative ∨
    ∨-assoc       : Associative ∨
    ∨-cong        : Congruent₂ ∨
    ∧-comm        : Commutative ∧
    ∧-assoc       : Associative ∧
    ∧-cong        : Congruent₂ ∧
    absorptive    : Absorptive ∨ ∧
  open IsEquivalence isEquivalence public
  ∨-absorbs-∧ : ∨ Absorbs ∧
  ∨-absorbs-∧ = proj₁ absorptive
  ∧-absorbs-∨ : ∧ Absorbs ∨
  ∧-absorbs-∨ = proj₂ absorptive
  ∧-congˡ : LeftCongruent ∧
  ∧-congˡ y≈z = ∧-cong refl y≈z
  ∧-congʳ : RightCongruent ∧
  ∧-congʳ y≈z = ∧-cong y≈z refl
  ∨-congˡ : LeftCongruent ∨
  ∨-congˡ y≈z = ∨-cong refl y≈z
  ∨-congʳ : RightCongruent ∨
  ∨-congʳ y≈z = ∨-cong y≈z refl
record IsDistributiveLattice (∨ ∧ : Op₂ A) : Set (a ⊔ ℓ) where
  field
    isLattice   : IsLattice ∨ ∧
    ∨-distrib-∧ : ∨ DistributesOver ∧
    ∧-distrib-∨ : ∧ DistributesOver ∨
  open IsLattice isLattice public
  ∨-distribˡ-∧ : ∨ DistributesOverˡ ∧
  ∨-distribˡ-∧ = proj₁ ∨-distrib-∧
  ∨-distribʳ-∧ : ∨ DistributesOverʳ ∧
  ∨-distribʳ-∧ = proj₂ ∨-distrib-∧
  ∧-distribˡ-∨ : ∧ DistributesOverˡ ∨
  ∧-distribˡ-∨ = proj₁ ∧-distrib-∨
  ∧-distribʳ-∨ : ∧ DistributesOverʳ ∨
  ∧-distribʳ-∨ = proj₂ ∧-distrib-∨
record IsBooleanAlgebra (∨ ∧ : Op₂ A) (¬ : Op₁ A) (⊤ ⊥ : A) : Set (a ⊔ ℓ)
  where
  field
    isDistributiveLattice : IsDistributiveLattice ∨ ∧
    ∨-complement          : Inverse ⊤ ¬ ∨
    ∧-complement          : Inverse ⊥ ¬ ∧
    ¬-cong                : Congruent₁ ¬
  open IsDistributiveLattice isDistributiveLattice public
  ∨-complementˡ : LeftInverse ⊤ ¬ ∨
  ∨-complementˡ = proj₁ ∨-complement
  ∨-complementʳ : RightInverse ⊤ ¬ ∨
  ∨-complementʳ = proj₂ ∨-complement
  ∧-complementˡ : LeftInverse ⊥ ¬ ∧
  ∧-complementˡ = proj₁ ∧-complement
  ∧-complementʳ : RightInverse ⊥ ¬ ∧
  ∧-complementʳ = proj₂ ∧-complement