| 
									
										
										
										
											2023-04-04 21:08:31 -07:00
										 |  |  |  | module Lattice where
 | 
					
						
							|  |  |  |  | 
 | 
					
						
							| 
									
										
										
										
											2023-09-02 20:36:12 -07:00
										 |  |  |  | open import Equivalence
 | 
					
						
							| 
									
										
										
										
											2023-09-03 21:05:57 -07:00
										 |  |  |  | import Chain
 | 
					
						
							| 
									
										
										
										
											2023-09-02 20:36:12 -07:00
										 |  |  |  | 
 | 
					
						
							| 
									
										
										
										
											2023-09-03 21:05:57 -07:00
										 |  |  |  | open import Relation.Binary.Core using (_Preserves_⟶_ )
 | 
					
						
							| 
									
										
										
										
											2023-09-23 16:39:11 -07:00
										 |  |  |  | open import Relation.Nullary using (Dec; ¬_)
 | 
					
						
							|  |  |  |  | open import Data.Nat as Nat using (ℕ)
 | 
					
						
							|  |  |  |  | open import Data.Product using (_×_; Σ; _,_)
 | 
					
						
							| 
									
										
										
										
											2023-07-13 21:50:27 -07:00
										 |  |  |  | open import Data.Sum using (_⊎_; inj₁; inj₂)
 | 
					
						
							| 
									
										
										
										
											2023-08-20 20:49:08 -07:00
										 |  |  |  | open import Agda.Primitive using (lsuc; Level) renaming (_⊔_ to _⊔ℓ_)
 | 
					
						
							|  |  |  |  | open import Function.Definitions using (Injective)
 | 
					
						
							| 
									
										
										
										
											2023-09-03 23:56:39 -07:00
										 |  |  |  | 
 | 
					
						
							| 
									
										
										
										
											2023-09-17 19:43:24 -07:00
										 |  |  |  | IsDecidable : ∀ {a} {A : Set a} (R : A → A → Set a) → Set a
 | 
					
						
							|  |  |  |  | IsDecidable {a} {A} R = ∀ (a₁ a₂ : A) → Dec (R a₁ a₂)
 | 
					
						
							| 
									
										
										
										
											2023-08-20 18:35:57 -07:00
										 |  |  |  | 
 | 
					
						
							| 
									
										
										
										
											2023-07-15 14:40:11 -07:00
										 |  |  |  | record IsSemilattice {a} (A : Set a)
 | 
					
						
							|  |  |  |  |     (_≈_ : A → A → Set a)
 | 
					
						
							|  |  |  |  |     (_⊔_ : A → A → A) : Set a where
 | 
					
						
							|  |  |  |  | 
 | 
					
						
							| 
									
										
										
										
											2023-08-20 19:57:26 -07:00
										 |  |  |  |     _≼_ : A → A → Set a
 | 
					
						
							|  |  |  |  |     a ≼ b = Σ A (λ c → (a ⊔ c) ≈ b)
 | 
					
						
							|  |  |  |  | 
 | 
					
						
							|  |  |  |  |     _≺_ : A → A → Set a
 | 
					
						
							|  |  |  |  |     a ≺ b = (a ≼ b) × (¬ a ≈ b)
 | 
					
						
							|  |  |  |  | 
 | 
					
						
							| 
									
										
										
										
											2023-04-06 23:08:49 -07:00
										 |  |  |  |     field
 | 
					
						
							| 
									
										
										
										
											2023-07-15 15:16:51 -07:00
										 |  |  |  |         ≈-equiv : IsEquivalence A _≈_
 | 
					
						
							| 
									
										
										
										
											2023-09-03 17:08:37 -07:00
										 |  |  |  |         ≈-⊔-cong : ∀ {a₁ a₂ a₃ a₄} → a₁ ≈ a₂ → a₃ ≈ a₄ → (a₁ ⊔ a₃) ≈ (a₂ ⊔ a₄)
 | 
					
						
							| 
									
										
										
										
											2023-07-15 15:16:51 -07:00
										 |  |  |  | 
 | 
					
						
							| 
									
										
										
										
											2023-07-15 14:40:11 -07:00
										 |  |  |  |         ⊔-assoc : (x y z : A) → ((x ⊔ y) ⊔ z) ≈ (x ⊔ (y ⊔ z))
 | 
					
						
							|  |  |  |  |         ⊔-comm : (x y : A) → (x ⊔ y) ≈ (y ⊔ x)
 | 
					
						
							|  |  |  |  |         ⊔-idemp : (x : A) → (x ⊔ x) ≈ x
 | 
					
						
							|  |  |  |  | 
 | 
					
						
							| 
									
										
										
										
											2023-09-03 19:33:04 -07:00
										 |  |  |  |     open IsEquivalence ≈-equiv public
 | 
					
						
							|  |  |  |  | 
 | 
					
						
							| 
									
										
										
										
											2023-08-20 21:53:27 -07:00
										 |  |  |  |     ≼-refl : ∀ (a : A) → a ≼ a
 | 
					
						
							|  |  |  |  |     ≼-refl a = (a , ⊔-idemp a)
 | 
					
						
							|  |  |  |  | 
 | 
					
						
							| 
									
										
										
										
											2023-09-03 19:33:04 -07:00
										 |  |  |  |     ≼-cong : ∀ {a₁ a₂ a₃ a₄ : A} → a₁ ≈ a₂ → a₃ ≈ a₄ → a₁ ≼ a₃ → a₂ ≼ a₄
 | 
					
						
							|  |  |  |  |     ≼-cong a₁≈a₂ a₃≈a₄ (c₁ , a₁⊔c₁≈a₃) = (c₁ , ≈-trans (≈-⊔-cong (≈-sym a₁≈a₂) ≈-refl) (≈-trans a₁⊔c₁≈a₃ a₃≈a₄))
 | 
					
						
							|  |  |  |  | 
 | 
					
						
							|  |  |  |  |     ≺-cong : ∀ {a₁ a₂ a₃ a₄ : A} → a₁ ≈ a₂ → a₃ ≈ a₄ → a₁ ≺ a₃ → a₂ ≺ a₄
 | 
					
						
							|  |  |  |  |     ≺-cong a₁≈a₂ a₃≈a₄ (a₁≼a₃ , a₁̷≈a₃) =
 | 
					
						
							|  |  |  |  |         ( ≼-cong a₁≈a₂ a₃≈a₄ a₁≼a₃
 | 
					
						
							|  |  |  |  |         , λ a₂≈a₄ → a₁̷≈a₃ (≈-trans a₁≈a₂ (≈-trans a₂≈a₄ (≈-sym a₃≈a₄)))
 | 
					
						
							|  |  |  |  |         )
 | 
					
						
							| 
									
										
										
										
											2023-07-15 15:16:51 -07:00
										 |  |  |  | 
 | 
					
						
							| 
									
										
										
										
											2023-07-15 14:40:11 -07:00
										 |  |  |  | record IsLattice {a} (A : Set a)
 | 
					
						
							|  |  |  |  |     (_≈_ : A → A → Set a)
 | 
					
						
							|  |  |  |  |     (_⊔_ : A → A → A)
 | 
					
						
							|  |  |  |  |     (_⊓_ : A → A → A) : Set a where
 | 
					
						
							| 
									
										
										
										
											2023-04-04 21:08:31 -07:00
										 |  |  |  | 
 | 
					
						
							|  |  |  |  |     field
 | 
					
						
							| 
									
										
										
										
											2023-07-15 14:40:11 -07:00
										 |  |  |  |         joinSemilattice : IsSemilattice A _≈_ _⊔_
 | 
					
						
							|  |  |  |  |         meetSemilattice : IsSemilattice A _≈_ _⊓_
 | 
					
						
							| 
									
										
										
										
											2023-04-06 23:08:49 -07:00
										 |  |  |  | 
 | 
					
						
							| 
									
										
										
										
											2023-07-15 14:40:11 -07:00
										 |  |  |  |         absorb-⊔-⊓ : (x y : A) → (x ⊔ (x ⊓ y)) ≈ x
 | 
					
						
							|  |  |  |  |         absorb-⊓-⊔ : (x y : A) → (x ⊓ (x ⊔ y)) ≈ x
 | 
					
						
							| 
									
										
										
										
											2023-04-04 21:08:31 -07:00
										 |  |  |  | 
 | 
					
						
							| 
									
										
										
										
											2023-04-06 23:08:49 -07:00
										 |  |  |  |     open IsSemilattice joinSemilattice public
 | 
					
						
							| 
									
										
										
										
											2023-08-20 19:57:26 -07:00
										 |  |  |  |     open IsSemilattice meetSemilattice public using () renaming
 | 
					
						
							| 
									
										
										
										
											2023-04-06 23:08:49 -07:00
										 |  |  |  |         ( ⊔-assoc to ⊓-assoc
 | 
					
						
							|  |  |  |  |         ; ⊔-comm to ⊓-comm
 | 
					
						
							|  |  |  |  |         ; ⊔-idemp to ⊓-idemp
 | 
					
						
							| 
									
										
										
										
											2023-09-03 17:08:37 -07:00
										 |  |  |  |         ; ≈-⊔-cong to ≈-⊓-cong
 | 
					
						
							| 
									
										
										
										
											2023-08-20 21:53:27 -07:00
										 |  |  |  |         ; _≼_ to _≽_
 | 
					
						
							|  |  |  |  |         ; _≺_ to _≻_
 | 
					
						
							|  |  |  |  |         ; ≼-refl to ≽-refl
 | 
					
						
							| 
									
										
										
										
											2023-04-06 23:08:49 -07:00
										 |  |  |  |         )
 | 
					
						
							|  |  |  |  | 
 | 
					
						
							| 
									
										
										
										
											2023-09-23 16:20:58 -07:00
										 |  |  |  |     FixedHeight : ∀ (h : ℕ) → Set a
 | 
					
						
							|  |  |  |  |     FixedHeight h = Chain.Height (_≈_) ≈-equiv _≺_ ≺-cong h
 | 
					
						
							|  |  |  |  | 
 | 
					
						
							| 
									
										
										
										
											2023-08-20 18:35:57 -07:00
										 |  |  |  | record IsFiniteHeightLattice {a} (A : Set a)
 | 
					
						
							|  |  |  |  |     (h : ℕ)
 | 
					
						
							|  |  |  |  |     (_≈_ : A → A → Set a)
 | 
					
						
							|  |  |  |  |     (_⊔_ : A → A → A)
 | 
					
						
							|  |  |  |  |     (_⊓_ : A → A → A) : Set (lsuc a) where
 | 
					
						
							|  |  |  |  | 
 | 
					
						
							|  |  |  |  |     field
 | 
					
						
							|  |  |  |  |         isLattice : IsLattice A _≈_ _⊔_ _⊓_
 | 
					
						
							|  |  |  |  | 
 | 
					
						
							|  |  |  |  |     open IsLattice isLattice public
 | 
					
						
							|  |  |  |  | 
 | 
					
						
							| 
									
										
										
										
											2023-09-23 16:20:58 -07:00
										 |  |  |  |     field
 | 
					
						
							|  |  |  |  |         fixedHeight : FixedHeight h
 | 
					
						
							| 
									
										
										
										
											2023-09-03 19:33:04 -07:00
										 |  |  |  | 
 | 
					
						
							| 
									
										
										
										
											2023-08-20 20:49:08 -07:00
										 |  |  |  | module _ {a b} {A : Set a} {B : Set b}
 | 
					
						
							| 
									
										
										
										
											2023-08-20 21:53:27 -07:00
										 |  |  |  |     (_≼₁_ : 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}
 | 
					
						
							| 
									
										
										
										
											2023-08-20 20:49:08 -07:00
										 |  |  |  |     (slA : IsSemilattice A _≈₁_ _⊔₁_) (slB : IsSemilattice B _≈₂_ _⊔₂_) where
 | 
					
						
							|  |  |  |  | 
 | 
					
						
							| 
									
										
										
										
											2023-09-03 21:05:57 -07:00
										 |  |  |  |     open IsSemilattice slA renaming (_≼_ to _≼₁_; _≺_ to _≺₁_; ≈-equiv to ≈₁-equiv; ≺-cong to ≺₁-cong)
 | 
					
						
							|  |  |  |  |     open IsSemilattice slB renaming (_≼_ to _≼₂_; _≺_ to _≺₂_; ≈-equiv to ≈₂-equiv; ≺-cong to ≺₂-cong)
 | 
					
						
							| 
									
										
										
										
											2023-08-20 20:49:08 -07:00
										 |  |  |  | 
 | 
					
						
							| 
									
										
										
										
											2023-09-03 21:05:57 -07:00
										 |  |  |  |     open Chain _≈₁_ ≈₁-equiv _≺₁_ ≺₁-cong using () renaming (Chain to Chain₁; step to step₁; done to done₁)
 | 
					
						
							|  |  |  |  |     open Chain _≈₂_ ≈₂-equiv _≺₂_ ≺₂-cong using () renaming (Chain to Chain₂; step to step₂; done to done₂)
 | 
					
						
							|  |  |  |  | 
 | 
					
						
							|  |  |  |  |     Chain-map : ∀ (f : A → B) →
 | 
					
						
							|  |  |  |  |                 Monotonic _≼₁_ _≼₂_ f →
 | 
					
						
							|  |  |  |  |                 Injective _≈₁_ _≈₂_ f →
 | 
					
						
							|  |  |  |  |                 f Preserves _≈₁_ ⟶  _≈₂_ →
 | 
					
						
							|  |  |  |  |                 ∀ {a₁ a₂ : A} {n : ℕ} → Chain₁ a₁ a₂ n → Chain₂ (f a₁) (f a₂) n
 | 
					
						
							|  |  |  |  |     Chain-map f Monotonicᶠ Injectiveᶠ Preservesᶠ (done₁ a₁≈a₂) =
 | 
					
						
							|  |  |  |  |         done₂ (Preservesᶠ a₁≈a₂)
 | 
					
						
							|  |  |  |  |     Chain-map f Monotonicᶠ Injectiveᶠ Preservesᶠ (step₁ (a₁≼₁a , a₁̷≈₁a) a≈₁a' a'a₂) =
 | 
					
						
							| 
									
										
										
										
											2023-08-20 20:49:08 -07:00
										 |  |  |  |         let fa₁≺₂fa = (Monotonicᶠ a₁≼₁a , λ fa₁≈₂fa → a₁̷≈₁a (Injectiveᶠ fa₁≈₂fa))
 | 
					
						
							| 
									
										
										
										
											2023-09-03 21:05:57 -07:00
										 |  |  |  |             fa≈fa' = Preservesᶠ a≈₁a'
 | 
					
						
							|  |  |  |  |         in step₂ fa₁≺₂fa fa≈fa' (Chain-map f Monotonicᶠ Injectiveᶠ Preservesᶠ a'a₂)
 | 
					
						
							| 
									
										
										
										
											2023-08-20 20:49:08 -07:00
										 |  |  |  | 
 | 
					
						
							| 
									
										
										
										
											2023-07-15 13:12:21 -07:00
										 |  |  |  | record Semilattice {a} (A : Set a) : Set (lsuc a) where
 | 
					
						
							|  |  |  |  |     field
 | 
					
						
							| 
									
										
										
										
											2023-07-15 14:40:11 -07:00
										 |  |  |  |         _≈_ : A → A → Set a
 | 
					
						
							| 
									
										
										
										
											2023-07-15 13:12:21 -07:00
										 |  |  |  |         _⊔_ : A → A → A
 | 
					
						
							|  |  |  |  | 
 | 
					
						
							| 
									
										
										
										
											2023-07-15 14:40:11 -07:00
										 |  |  |  |         isSemilattice : IsSemilattice A _≈_ _⊔_
 | 
					
						
							| 
									
										
										
										
											2023-07-15 13:12:21 -07:00
										 |  |  |  | 
 | 
					
						
							|  |  |  |  |     open IsSemilattice isSemilattice public
 | 
					
						
							|  |  |  |  | 
 | 
					
						
							| 
									
										
										
										
											2023-04-06 23:08:49 -07:00
										 |  |  |  | 
 | 
					
						
							|  |  |  |  | record Lattice {a} (A : Set a) : Set (lsuc a) where
 | 
					
						
							|  |  |  |  |     field
 | 
					
						
							| 
									
										
										
										
											2023-07-15 14:40:11 -07:00
										 |  |  |  |         _≈_ : A → A → Set a
 | 
					
						
							| 
									
										
										
										
											2023-04-06 23:08:49 -07:00
										 |  |  |  |         _⊔_ : A → A → A
 | 
					
						
							|  |  |  |  |         _⊓_ : A → A → A
 | 
					
						
							|  |  |  |  | 
 | 
					
						
							| 
									
										
										
										
											2023-07-15 14:40:11 -07:00
										 |  |  |  |         isLattice : IsLattice A _≈_ _⊔_ _⊓_
 | 
					
						
							| 
									
										
										
										
											2023-04-06 23:08:49 -07:00
										 |  |  |  | 
 | 
					
						
							|  |  |  |  |     open IsLattice isLattice public
 |