Prove monotonicity of lub in one argument
Signed-off-by: Danila Fedorin <danila.fedorin@gmail.com>
This commit is contained in:
parent
ae3e2c28b0
commit
65d1590358
32
Lattice.agda
32
Lattice.agda
|
@ -14,6 +14,12 @@ open import Function.Definitions using (Injective)
|
|||
IsDecidable : ∀ {a} {A : Set a} (R : A → A → Set a) → Set a
|
||||
IsDecidable {a} {A} R = ∀ (a₁ a₂ : A) → Dec (R a₁ a₂)
|
||||
|
||||
module _ {a b} {A : Set a} {B : Set b}
|
||||
(_≼₁_ : A → A → Set a) (_≼₂_ : B → B → Set b) where
|
||||
|
||||
Monotonic : (A → B) → Set (a ⊔ℓ b)
|
||||
Monotonic f = ∀ {a₁ a₂ : A} → a₁ ≼₁ a₂ → f a₁ ≼₂ f a₂
|
||||
|
||||
record IsSemilattice {a} (A : Set a)
|
||||
(_≈_ : A → A → Set a)
|
||||
(_⊔_ : A → A → A) : Set a where
|
||||
|
@ -36,6 +42,26 @@ record IsSemilattice {a} (A : Set a)
|
|||
|
||||
open import Relation.Binary.Reasoning.Base.Single _≈_ ≈-refl ≈-trans
|
||||
|
||||
⊔-Monotonicˡ : ∀ (a₁ : A) → Monotonic _≼_ _≼_ (λ a₂ → a₁ ⊔ a₂)
|
||||
⊔-Monotonicˡ a {a₁} {a₂} a₁≼a₂ = ≈-trans (≈-sym lhs) (≈-⊔-cong (≈-refl {a}) a₁≼a₂)
|
||||
where
|
||||
lhs =
|
||||
begin
|
||||
a ⊔ (a₁ ⊔ a₂)
|
||||
∼⟨ ≈-⊔-cong (≈-sym (⊔-idemp _)) ≈-refl ⟩
|
||||
(a ⊔ a) ⊔ (a₁ ⊔ a₂)
|
||||
∼⟨ ⊔-assoc _ _ _ ⟩
|
||||
a ⊔ (a ⊔ (a₁ ⊔ a₂))
|
||||
∼⟨ ≈-⊔-cong ≈-refl (≈-sym (⊔-assoc _ _ _)) ⟩
|
||||
a ⊔ ((a ⊔ a₁) ⊔ a₂)
|
||||
∼⟨ ≈-⊔-cong ≈-refl (≈-⊔-cong (⊔-comm _ _) ≈-refl) ⟩
|
||||
a ⊔ ((a₁ ⊔ a) ⊔ a₂)
|
||||
∼⟨ ≈-⊔-cong ≈-refl (⊔-assoc _ _ _) ⟩
|
||||
a ⊔ (a₁ ⊔ (a ⊔ a₂))
|
||||
∼⟨ ≈-sym (⊔-assoc _ _ _) ⟩
|
||||
(a ⊔ a₁) ⊔ (a ⊔ a₂)
|
||||
∎
|
||||
|
||||
≼-refl : ∀ (a : A) → a ≼ a
|
||||
≼-refl a = ⊔-idemp a
|
||||
|
||||
|
@ -97,12 +123,6 @@ record IsFiniteHeightLattice {a} (A : Set a)
|
|||
field
|
||||
fixedHeight : FixedHeight h
|
||||
|
||||
module _ {a b} {A : Set a} {B : Set b}
|
||||
(_≼₁_ : A → A → Set a) (_≼₂_ : B → B → Set b) where
|
||||
|
||||
Monotonic : (A → B) → Set (a ⊔ℓ b)
|
||||
Monotonic f = ∀ {a₁ a₂ : A} → a₁ ≼₁ a₂ → f a₁ ≼₂ f a₂
|
||||
|
||||
module ChainMapping {a b} {A : Set a} {B : Set b}
|
||||
{_≈₁_ : A → A → Set a} {_≈₂_ : B → B → Set b}
|
||||
{_⊔₁_ : A → A → A} {_⊔₂_ : B → B → B}
|
||||
|
|
Loading…
Reference in New Issue
Block a user