2024-02-11 14:18:58 -08:00
|
|
|
|
open import Lattice
|
|
|
|
|
|
2024-02-19 12:36:46 -08:00
|
|
|
|
-- Due to universe levels, it becomes relatively annoying to handle the case
|
|
|
|
|
-- where the levels of A and B are not the same. For the time being, constrain
|
|
|
|
|
-- them to be the same.
|
|
|
|
|
|
2024-02-11 14:18:58 -08:00
|
|
|
|
module Lattice.IterProd {a} {A B : Set a}
|
|
|
|
|
(_≈₁_ : A → A → Set a) (_≈₂_ : B → B → Set a)
|
|
|
|
|
(_⊔₁_ : A → A → A) (_⊔₂_ : B → B → B)
|
|
|
|
|
(_⊓₁_ : A → A → A) (_⊓₂_ : B → B → B)
|
|
|
|
|
(lA : IsLattice A _≈₁_ _⊔₁_ _⊓₁_) (lB : IsLattice B _≈₂_ _⊔₂_ _⊓₂_) where
|
|
|
|
|
|
2024-02-11 15:36:12 -08:00
|
|
|
|
open import Agda.Primitive using (lsuc)
|
2024-05-09 19:45:15 -07:00
|
|
|
|
open import Data.Nat using (ℕ; zero; suc; _+_)
|
|
|
|
|
open import Data.Product using (_×_; _,_; proj₁; proj₂)
|
2024-05-09 20:20:11 -07:00
|
|
|
|
open import Relation.Binary.PropositionalEquality as Eq using (_≡_; refl; cong)
|
2024-02-11 14:18:58 -08:00
|
|
|
|
open import Utils using (iterate)
|
2024-05-09 20:11:04 -07:00
|
|
|
|
open import Chain using (Height)
|
2024-02-11 14:18:58 -08:00
|
|
|
|
|
2024-02-11 15:36:12 -08:00
|
|
|
|
open IsLattice lA renaming (FixedHeight to FixedHeight₁)
|
|
|
|
|
open IsLattice lB renaming (FixedHeight to FixedHeight₂)
|
|
|
|
|
|
2024-02-11 14:18:58 -08:00
|
|
|
|
IterProd : ℕ → Set a
|
|
|
|
|
IterProd k = iterate k (λ t → A × t) B
|
|
|
|
|
|
|
|
|
|
-- To make iteration more convenient, package the definitions in Lattice
|
|
|
|
|
-- records, perform the recursion, and unpackage.
|
2024-03-01 21:02:56 -08:00
|
|
|
|
--
|
|
|
|
|
|
|
|
|
|
-- If we prove isLattice and isFiniteHeightLattice by induction separately,
|
|
|
|
|
-- we lose the connection between the operations (which should be the same)
|
|
|
|
|
-- that are built up by the two iterations. So, do everything in one iteration.
|
|
|
|
|
-- This requires some odd code.
|
|
|
|
|
|
2024-05-09 19:45:15 -07:00
|
|
|
|
build : A → B → (k : ℕ) → IterProd k
|
|
|
|
|
build _ b zero = b
|
|
|
|
|
build a b (suc s) = (a , build a b s)
|
|
|
|
|
|
2024-03-01 21:02:56 -08:00
|
|
|
|
private
|
|
|
|
|
record RequiredForFixedHeight : Set (lsuc a) where
|
2024-05-09 19:45:15 -07:00
|
|
|
|
field
|
2024-03-01 21:02:56 -08:00
|
|
|
|
≈₁-dec : IsDecidable _≈₁_
|
|
|
|
|
≈₂-dec : IsDecidable _≈₂_
|
|
|
|
|
h₁ h₂ : ℕ
|
|
|
|
|
fhA : FixedHeight₁ h₁
|
|
|
|
|
fhB : FixedHeight₂ h₂
|
|
|
|
|
|
2024-05-09 19:45:15 -07:00
|
|
|
|
⊥₁ : A
|
2024-05-09 20:11:04 -07:00
|
|
|
|
⊥₁ = Height.⊥ fhA
|
2024-05-09 19:45:15 -07:00
|
|
|
|
|
|
|
|
|
⊥₂ : B
|
2024-05-09 20:11:04 -07:00
|
|
|
|
⊥₂ = Height.⊥ fhB
|
2024-05-09 19:45:15 -07:00
|
|
|
|
|
|
|
|
|
⊥k : ∀ (k : ℕ) → IterProd k
|
|
|
|
|
⊥k = build ⊥₁ ⊥₂
|
|
|
|
|
|
|
|
|
|
record IsFiniteHeightWithBotAndDecEq {A : Set a} {_≈_ : A → A → Set a} {_⊔_ : A → A → A} {_⊓_ : A → A → A} (isLattice : IsLattice A _≈_ _⊔_ _⊓_) (⊥ : A) : Set (lsuc a) where
|
2024-03-01 21:02:56 -08:00
|
|
|
|
field
|
|
|
|
|
height : ℕ
|
|
|
|
|
fixedHeight : IsLattice.FixedHeight isLattice height
|
|
|
|
|
≈-dec : IsDecidable _≈_
|
|
|
|
|
|
2024-05-09 20:20:11 -07:00
|
|
|
|
⊥-correct : Height.⊥ fixedHeight ≡ ⊥
|
2024-05-09 19:45:15 -07:00
|
|
|
|
|
|
|
|
|
record Everything (k : ℕ) : Set (lsuc a) where
|
|
|
|
|
T = IterProd k
|
|
|
|
|
|
2024-03-01 21:02:56 -08:00
|
|
|
|
field
|
2024-05-09 19:45:15 -07:00
|
|
|
|
_≈_ : T → T → Set a
|
|
|
|
|
_⊔_ : T → T → T
|
|
|
|
|
_⊓_ : T → T → T
|
2024-03-01 21:02:56 -08:00
|
|
|
|
|
2024-05-09 19:45:15 -07:00
|
|
|
|
isLattice : IsLattice T _≈_ _⊔_ _⊓_
|
|
|
|
|
isFiniteHeightIfSupported :
|
|
|
|
|
∀ (req : RequiredForFixedHeight) →
|
|
|
|
|
IsFiniteHeightWithBotAndDecEq isLattice (RequiredForFixedHeight.⊥k req k)
|
2024-03-01 21:02:56 -08:00
|
|
|
|
|
2024-05-09 19:45:15 -07:00
|
|
|
|
everything : ∀ (k : ℕ) → Everything k
|
2024-03-01 21:02:56 -08:00
|
|
|
|
everything 0 = record
|
2024-02-11 14:18:58 -08:00
|
|
|
|
{ _≈_ = _≈₂_
|
|
|
|
|
; _⊔_ = _⊔₂_
|
|
|
|
|
; _⊓_ = _⊓₂_
|
|
|
|
|
; isLattice = lB
|
2024-03-01 21:02:56 -08:00
|
|
|
|
; isFiniteHeightIfSupported = λ req → record
|
|
|
|
|
{ height = RequiredForFixedHeight.h₂ req
|
|
|
|
|
; fixedHeight = RequiredForFixedHeight.fhB req
|
|
|
|
|
; ≈-dec = RequiredForFixedHeight.≈₂-dec req
|
2024-05-09 20:20:11 -07:00
|
|
|
|
; ⊥-correct = refl
|
2024-03-01 21:02:56 -08:00
|
|
|
|
}
|
2024-02-11 14:18:58 -08:00
|
|
|
|
}
|
2024-03-01 21:02:56 -08:00
|
|
|
|
everything (suc k') = record
|
|
|
|
|
{ _≈_ = P._≈_
|
|
|
|
|
; _⊔_ = P._⊔_
|
|
|
|
|
; _⊓_ = P._⊓_
|
|
|
|
|
; isLattice = P.isLattice
|
|
|
|
|
; isFiniteHeightIfSupported = λ req →
|
|
|
|
|
let
|
|
|
|
|
fhlRest = Everything.isFiniteHeightIfSupported everythingRest req
|
|
|
|
|
in
|
|
|
|
|
record
|
2024-05-09 19:45:15 -07:00
|
|
|
|
{ height = (RequiredForFixedHeight.h₁ req) + IsFiniteHeightWithBotAndDecEq.height fhlRest
|
2024-03-01 21:02:56 -08:00
|
|
|
|
; fixedHeight =
|
|
|
|
|
P.fixedHeight
|
2024-05-09 19:45:15 -07:00
|
|
|
|
(RequiredForFixedHeight.≈₁-dec req) (IsFiniteHeightWithBotAndDecEq.≈-dec fhlRest)
|
|
|
|
|
(RequiredForFixedHeight.h₁ req) (IsFiniteHeightWithBotAndDecEq.height fhlRest)
|
|
|
|
|
(RequiredForFixedHeight.fhA req) (IsFiniteHeightWithBotAndDecEq.fixedHeight fhlRest)
|
|
|
|
|
; ≈-dec = P.≈-dec (RequiredForFixedHeight.≈₁-dec req) (IsFiniteHeightWithBotAndDecEq.≈-dec fhlRest)
|
2024-05-09 20:20:11 -07:00
|
|
|
|
; ⊥-correct =
|
|
|
|
|
cong ((Height.⊥ (RequiredForFixedHeight.fhA req)) ,_)
|
|
|
|
|
(IsFiniteHeightWithBotAndDecEq.⊥-correct fhlRest)
|
2024-03-01 21:02:56 -08:00
|
|
|
|
}
|
2024-02-11 14:18:58 -08:00
|
|
|
|
}
|
|
|
|
|
where
|
2024-03-01 21:02:56 -08:00
|
|
|
|
everythingRest = everything k'
|
2024-02-11 14:18:58 -08:00
|
|
|
|
|
2024-03-01 21:02:56 -08:00
|
|
|
|
import Lattice.Prod
|
|
|
|
|
_≈₁_ (Everything._≈_ everythingRest)
|
|
|
|
|
_⊔₁_ (Everything._⊔_ everythingRest)
|
|
|
|
|
_⊓₁_ (Everything._⊓_ everythingRest)
|
|
|
|
|
lA (Everything.isLattice everythingRest) as P
|
2024-02-11 15:36:12 -08:00
|
|
|
|
|
2024-02-11 20:45:14 -08:00
|
|
|
|
module _ (k : ℕ) where
|
2024-03-01 21:02:56 -08:00
|
|
|
|
open Everything (everything k) using (_≈_; _⊔_; _⊓_; isLattice) public
|
|
|
|
|
open Lattice.IsLattice isLattice public
|
2024-02-11 15:36:12 -08:00
|
|
|
|
|
2024-03-01 21:12:22 -08:00
|
|
|
|
lattice : Lattice (IterProd k)
|
|
|
|
|
lattice = record
|
|
|
|
|
{ _≈_ = _≈_
|
|
|
|
|
; _⊔_ = _⊔_
|
|
|
|
|
; _⊓_ = _⊓_
|
|
|
|
|
; isLattice = isLattice
|
|
|
|
|
}
|
|
|
|
|
|
2024-03-01 21:02:56 -08:00
|
|
|
|
module _ (≈₁-dec : IsDecidable _≈₁_) (≈₂-dec : IsDecidable _≈₂_)
|
|
|
|
|
(h₁ h₂ : ℕ)
|
|
|
|
|
(fhA : FixedHeight₁ h₁) (fhB : FixedHeight₂ h₂) where
|
2024-02-11 14:18:58 -08:00
|
|
|
|
|
2024-02-11 20:45:14 -08:00
|
|
|
|
private
|
2024-03-01 21:02:56 -08:00
|
|
|
|
required : RequiredForFixedHeight
|
|
|
|
|
required = record
|
|
|
|
|
{ ≈₁-dec = ≈₁-dec
|
|
|
|
|
; ≈₂-dec = ≈₂-dec
|
|
|
|
|
; h₁ = h₁
|
|
|
|
|
; h₂ = h₂
|
|
|
|
|
; fhA = fhA
|
|
|
|
|
; fhB = fhB
|
|
|
|
|
}
|
|
|
|
|
|
2024-05-09 19:45:15 -07:00
|
|
|
|
fixedHeight = IsFiniteHeightWithBotAndDecEq.fixedHeight (Everything.isFiniteHeightIfSupported (everything k) required)
|
|
|
|
|
|
2024-03-01 21:02:56 -08:00
|
|
|
|
isFiniteHeightLattice = record
|
|
|
|
|
{ isLattice = isLattice
|
2024-05-09 19:45:15 -07:00
|
|
|
|
; fixedHeight = fixedHeight
|
2024-02-11 20:45:14 -08:00
|
|
|
|
}
|
2024-03-01 21:12:22 -08:00
|
|
|
|
|
|
|
|
|
finiteHeightLattice : FiniteHeightLattice (IterProd k)
|
|
|
|
|
finiteHeightLattice = record
|
2024-05-09 19:45:15 -07:00
|
|
|
|
{ height = IsFiniteHeightWithBotAndDecEq.height (Everything.isFiniteHeightIfSupported (everything k) required)
|
2024-03-01 21:12:22 -08:00
|
|
|
|
; _≈_ = _≈_
|
|
|
|
|
; _⊔_ = _⊔_
|
|
|
|
|
; _⊓_ = _⊓_
|
|
|
|
|
; isFiniteHeightLattice = isFiniteHeightLattice
|
|
|
|
|
}
|
2024-05-09 20:20:11 -07:00
|
|
|
|
|
|
|
|
|
⊥-built : Height.⊥ fixedHeight ≡ (build (Height.⊥ fhA) (Height.⊥ fhB) k)
|
|
|
|
|
⊥-built = IsFiniteHeightWithBotAndDecEq.⊥-correct (Everything.isFiniteHeightIfSupported (everything k) required)
|
|
|
|
|
|