Re-export members of isLattice together with the record where needed

Signed-off-by: Danila Fedorin <danila.fedorin@gmail.com>
This commit is contained in:
Danila Fedorin 2025-01-04 22:43:13 -08:00
parent c0238fea25
commit ca375976b7
5 changed files with 6 additions and 11 deletions

View File

@ -55,12 +55,7 @@ open VariableValuesFiniteMap
; ∈k-dec to ∈k-decᵛ ; ∈k-dec to ∈k-decᵛ
; all-equal-keys to all-equal-keysᵛ ; all-equal-keys to all-equal-keysᵛ
; Provenance-union to Provenance-unionᵛ ; Provenance-union to Provenance-unionᵛ
) ; ⊔-Monotonicˡ to ⊔ᵛ-Monotonicˡ
public
open IsLattice isLatticeᵛ
using ()
renaming
( ⊔-Monotonicˡ to ⊔ᵛ-Monotonicˡ
; ⊔-Monotonicʳ to ⊔ᵛ-Monotonicʳ ; ⊔-Monotonicʳ to ⊔ᵛ-Monotonicʳ
; ⊔-idemp to ⊔ᵛ-idemp ; ⊔-idemp to ⊔ᵛ-idemp
) )

View File

@ -78,10 +78,9 @@ open AB.Plain 0ˢ using ()
; _≼_ to _≼ᵍ_ ; _≼_ to _≼ᵍ_
; _⊔_ to _⊔ᵍ_ ; _⊔_ to _⊔ᵍ_
; _⊓_ to _⊓ᵍ_ ; _⊓_ to _⊓ᵍ_
; ≼-trans to ≼ᵍ-trans
) )
open IsLattice isLatticeᵍ using () renaming (≼-trans to ≼ᵍ-trans)
plus : SignLattice SignLattice SignLattice plus : SignLattice SignLattice SignLattice
plus ⊥ᵍ _ = ⊥ᵍ plus ⊥ᵍ _ = ⊥ᵍ
plus _ ⊥ᵍ = ⊥ᵍ plus _ ⊥ᵍ = ⊥ᵍ

View File

@ -321,7 +321,7 @@ module Plain (x : A) where
; isLattice = isLattice ; isLattice = isLattice
} }
open IsLattice isLattice using (_≼_; _≺_; ⊔-Monotonicˡ; ⊔-Monotonicʳ) public open IsLattice isLattice using (_≼_; _≺_; ≼-trans; ⊔-Monotonicˡ; ⊔-Monotonicʳ) public
⊥≺[x] : (x : A) [ x ] ⊥≺[x] : (x : A) [ x ]
⊥≺[x] x = (≈-refl , λ ()) ⊥≺[x] x = (≈-refl , λ ())

View File

@ -184,7 +184,7 @@ private module WithKeys (ks : List A) where
; isLattice = isLattice ; isLattice = isLattice
} }
open IsLattice isLattice using (_≼_; ⊔-Monotonicˡ; ⊔-Monotonicʳ) public open IsLattice isLattice using (_≼_; ⊔-idemp; ⊔-Monotonicˡ; ⊔-Monotonicʳ) public
m₁≼m₂⇒m₁[k]≼m₂[k] : (fm₁ fm₂ : FiniteMap) {k : A} {v₁ v₂ : B} m₁≼m₂⇒m₁[k]≼m₂[k] : (fm₁ fm₂ : FiniteMap) {k : A} {v₁ v₂ : B}
fm₁ fm₂ (k , v₁) fm₁ (k , v₂) fm₂ v₁ ≼₂ v₂ fm₁ fm₂ (k , v₁) fm₁ (k , v₂) fm₂ v₁ ≼₂ v₂

View File

@ -106,6 +106,8 @@ instance
; isLattice = isLattice ; isLattice = isLattice
} }
open IsLattice isLattice using (_≼_; _≺_; ≺-cong) public
module _ {{≈₁-Decidable : IsDecidable _≈₁_}} {{≈₂-Decidable : IsDecidable _≈₂_}} where module _ {{≈₁-Decidable : IsDecidable _≈₁_}} {{≈₂-Decidable : IsDecidable _≈₂_}} where
open IsDecidable ≈₁-Decidable using () renaming (R-dec to ≈₁-dec) open IsDecidable ≈₁-Decidable using () renaming (R-dec to ≈₁-dec)
open IsDecidable ≈₂-Decidable using () renaming (R-dec to ≈₂-dec) open IsDecidable ≈₂-Decidable using () renaming (R-dec to ≈₂-dec)
@ -125,7 +127,6 @@ module _ {{≈₁-Decidable : IsDecidable _≈₁_}} {{≈₂-Decidable : IsDeci
{{fhA : FixedHeight₁ h₁}} {{fhB : FixedHeight₂ h₂}} where {{fhA : FixedHeight₁ h₁}} {{fhB : FixedHeight₂ h₂}} where
open import Data.Nat.Properties open import Data.Nat.Properties
open IsLattice isLattice using (_≼_; _≺_; ≺-cong)
module ChainMapping = ChainMapping joinSemilattice₁ isJoinSemilattice module ChainMapping = ChainMapping joinSemilattice₁ isJoinSemilattice
module ChainMapping = ChainMapping joinSemilattice₂ isJoinSemilattice module ChainMapping = ChainMapping joinSemilattice₂ isJoinSemilattice