2023-09-16 13:39:35 -07:00
|
|
|
|
open import Data.Nat as Nat using (ℕ; suc; _+_; _≤_)
|
|
|
|
|
|
open import Lattice
|
|
|
|
|
|
|
|
|
|
|
|
module Fixedpoint {a} {A : Set a}
|
|
|
|
|
|
{h : ℕ}
|
|
|
|
|
|
{_≈_ : A → A → Set a}
|
|
|
|
|
|
{_⊔_ : A → A → A} {_⊓_ : A → A → A}
|
2025-07-26 04:16:22 -07:00
|
|
|
|
{{ ≈-Decidable : IsDecidable _≈_ }}
|
|
|
|
|
|
{{flA : IsFiniteHeightLattice A h _≈_ _⊔_ _⊓_}}
|
2023-09-16 13:39:35 -07:00
|
|
|
|
(f : A → A)
|
|
|
|
|
|
(Monotonicᶠ : Monotonic (IsFiniteHeightLattice._≼_ flA)
|
|
|
|
|
|
(IsFiniteHeightLattice._≼_ flA) f) where
|
|
|
|
|
|
|
|
|
|
|
|
open import Data.Nat.Properties using (+-suc; +-comm)
|
|
|
|
|
|
open import Data.Product using (_×_; Σ; _,_; proj₁; proj₂)
|
2023-11-23 14:15:40 -08:00
|
|
|
|
open import Data.Empty using (⊥-elim)
|
2023-09-16 13:39:35 -07:00
|
|
|
|
open import Relation.Binary.PropositionalEquality using (_≡_; sym)
|
|
|
|
|
|
open import Relation.Nullary using (Dec; ¬_; yes; no)
|
|
|
|
|
|
|
2025-01-04 18:58:56 -08:00
|
|
|
|
open IsDecidable ≈-Decidable using () renaming (R-dec to ≈-dec)
|
2023-09-17 19:43:24 -07:00
|
|
|
|
open IsFiniteHeightLattice flA
|
|
|
|
|
|
|
2023-09-16 13:39:35 -07:00
|
|
|
|
import Chain
|
|
|
|
|
|
module ChainA = Chain _≈_ ≈-equiv _≺_ ≺-cong
|
|
|
|
|
|
|
|
|
|
|
|
private
|
2024-05-09 20:11:04 -07:00
|
|
|
|
open ChainA.Height fixedHeight
|
|
|
|
|
|
using ()
|
|
|
|
|
|
renaming
|
|
|
|
|
|
( ⊥ to ⊥ᴬ
|
|
|
|
|
|
; bounded to boundedᴬ
|
|
|
|
|
|
)
|
|
|
|
|
|
|
2023-09-16 13:39:35 -07:00
|
|
|
|
-- using 'g', for gas, here helps make sure the function terminates.
|
|
|
|
|
|
-- since A forms a fixed-height lattice, we must find a solution after
|
|
|
|
|
|
-- 'h' steps at most. Gas is set up such that as soon as it runs
|
|
|
|
|
|
-- out, we have exceeded h steps, which shouldn't be possible.
|
|
|
|
|
|
|
|
|
|
|
|
doStep : ∀ (g hᶜ : ℕ) (a₁ a₂ : A) (c : ChainA.Chain a₁ a₂ hᶜ) (g+hᶜ≡h : g + hᶜ ≡ suc h) (a₂≼fa₂ : a₂ ≼ f a₂) → Σ A (λ a → a ≈ f a)
|
2024-05-09 20:11:04 -07:00
|
|
|
|
doStep 0 hᶜ a₁ a₂ c g+hᶜ≡sh a₂≼fa₂ rewrite g+hᶜ≡sh = ⊥-elim (ChainA.Bounded-suc-n boundedᴬ c)
|
2023-09-16 13:39:35 -07:00
|
|
|
|
doStep (suc g') hᶜ a₁ a₂ c g+hᶜ≡sh a₂≼fa₂ rewrite sym (+-suc g' hᶜ)
|
|
|
|
|
|
with ≈-dec a₂ (f a₂)
|
|
|
|
|
|
... | yes a₂≈fa₂ = (a₂ , a₂≈fa₂)
|
|
|
|
|
|
... | no a₂̷≈fa₂ = doStep g' (suc hᶜ) a₁ (f a₂) c' g+hᶜ≡sh (Monotonicᶠ a₂≼fa₂)
|
|
|
|
|
|
where
|
|
|
|
|
|
a₂≺fa₂ : a₂ ≺ f a₂
|
|
|
|
|
|
a₂≺fa₂ = (a₂≼fa₂ , a₂̷≈fa₂)
|
|
|
|
|
|
|
|
|
|
|
|
c' : ChainA.Chain a₁ (f a₂) (suc hᶜ)
|
|
|
|
|
|
c' rewrite +-comm 1 hᶜ = ChainA.concat c (ChainA.step a₂≺fa₂ ≈-refl (ChainA.done (≈-refl {f a₂})))
|
|
|
|
|
|
|
|
|
|
|
|
fix : Σ A (λ a → a ≈ f a)
|
2025-07-26 04:16:22 -07:00
|
|
|
|
fix = doStep (suc h) 0 ⊥ᴬ ⊥ᴬ (ChainA.done ≈-refl) (+-comm (suc h) 0) (⊥≼ (f ⊥ᴬ))
|
2023-09-16 13:39:35 -07:00
|
|
|
|
|
|
|
|
|
|
aᶠ : A
|
|
|
|
|
|
aᶠ = proj₁ fix
|
|
|
|
|
|
|
|
|
|
|
|
aᶠ≈faᶠ : aᶠ ≈ f aᶠ
|
|
|
|
|
|
aᶠ≈faᶠ = proj₂ fix
|
|
|
|
|
|
|
|
|
|
|
|
private
|
2024-08-18 13:28:45 -07:00
|
|
|
|
stepPreservesLess : ∀ (g hᶜ : ℕ) (a₁ a₂ b : A) (b≈fb : b ≈ f b) (a₂≼a : a₂ ≼ b)
|
2023-09-16 13:39:35 -07:00
|
|
|
|
(c : ChainA.Chain a₁ a₂ hᶜ) (g+hᶜ≡h : g + hᶜ ≡ suc h)
|
|
|
|
|
|
(a₂≼fa₂ : a₂ ≼ f a₂) →
|
2024-08-18 13:28:45 -07:00
|
|
|
|
proj₁ (doStep g hᶜ a₁ a₂ c g+hᶜ≡h a₂≼fa₂) ≼ b
|
2024-05-09 20:11:04 -07:00
|
|
|
|
stepPreservesLess 0 _ _ _ _ _ _ c g+hᶜ≡sh _ rewrite g+hᶜ≡sh = ⊥-elim (ChainA.Bounded-suc-n boundedᴬ c)
|
2024-08-18 13:28:45 -07:00
|
|
|
|
stepPreservesLess (suc g') hᶜ a₁ a₂ b b≈fb a₂≼b c g+hᶜ≡sh a₂≼fa₂ rewrite sym (+-suc g' hᶜ)
|
2023-09-16 13:39:35 -07:00
|
|
|
|
with ≈-dec a₂ (f a₂)
|
2024-08-18 13:28:45 -07:00
|
|
|
|
... | yes _ = a₂≼b
|
|
|
|
|
|
... | no _ = stepPreservesLess g' _ _ _ b b≈fb (≼-cong ≈-refl (≈-sym b≈fb) (Monotonicᶠ a₂≼b)) _ _ _
|
2023-09-16 13:39:35 -07:00
|
|
|
|
|
|
|
|
|
|
aᶠ≼ : ∀ (a : A) → a ≈ f a → aᶠ ≼ a
|
2025-07-26 04:16:22 -07:00
|
|
|
|
aᶠ≼ a a≈fa = stepPreservesLess (suc h) 0 ⊥ᴬ ⊥ᴬ a a≈fa (⊥≼ a) (ChainA.done ≈-refl) (+-comm (suc h) 0) (⊥≼ (f ⊥ᴬ))
|