Prove that it's a least fixed point
Signed-off-by: Danila Fedorin <danila.fedorin@gmail.com>
This commit is contained in:
parent
c338fa3ee5
commit
cbebe599b2
27
Lattice.agda
27
Lattice.agda
|
@ -498,7 +498,7 @@ module FixedPoint {a} {A : Set a}
|
||||||
|
|
||||||
-- using 'g', for gas, here helps make sure the function terminates.
|
-- using 'g', for gas, here helps make sure the function terminates.
|
||||||
-- since A forms a fixed-height lattice, we must find a solution after
|
-- since A forms a fixed-height lattice, we must find a solution after
|
||||||
-- enough 'h' steps at most. Gas is set up such that as soon as it runs
|
-- '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.
|
-- 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)
|
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)
|
||||||
|
@ -514,5 +514,26 @@ module FixedPoint {a} {A : Set a}
|
||||||
c' : ChainA.Chain a₁ (f a₂) (suc hᶜ)
|
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₂})))
|
c' rewrite +-comm 1 hᶜ = ChainA.concat c (ChainA.step a₂≺fa₂ ≈-refl (ChainA.done (≈-refl {f a₂})))
|
||||||
|
|
||||||
fix : Σ A (λ a → a ≈ f a)
|
fix : Σ A (λ a → a ≈ f a)
|
||||||
fix = doStep (suc h) 0 ⊥ᴬ ⊥ᴬ (ChainA.done ≈-refl) (+-comm (suc h) 0) (⊥ᴬ≼ (f ⊥ᴬ))
|
fix = doStep (suc h) 0 ⊥ᴬ ⊥ᴬ (ChainA.done ≈-refl) (+-comm (suc h) 0) (⊥ᴬ≼ (f ⊥ᴬ))
|
||||||
|
|
||||||
|
aᶠ : A
|
||||||
|
aᶠ = proj₁ fix
|
||||||
|
|
||||||
|
aᶠ≈faᶠ : aᶠ ≈ f aᶠ
|
||||||
|
aᶠ≈faᶠ = proj₂ fix
|
||||||
|
|
||||||
|
private
|
||||||
|
stepPreservesLess : ∀ (g hᶜ : ℕ) (a₁ a₂ a : A) (a≈fa : a ≈ f a) (a₂≼a : a₂ ≼ a)
|
||||||
|
(c : ChainA.Chain a₁ a₂ hᶜ) (g+hᶜ≡h : g + hᶜ ≡ suc h)
|
||||||
|
(a₂≼fa₂ : a₂ ≼ f a₂) →
|
||||||
|
proj₁ (doStep g hᶜ a₁ a₂ c g+hᶜ≡h a₂≼fa₂) ≼ a
|
||||||
|
stepPreservesLess 0 _ _ _ _ _ _ c g+hᶜ≡sh _ rewrite g+hᶜ≡sh = absurd (ChainA.Bounded-suc-n (proj₂ fixedHeight) c)
|
||||||
|
stepPreservesLess (suc g') hᶜ a₁ a₂ a a≈fa a₂≼a c g+hᶜ≡sh a₂≼fa₂ rewrite sym (+-suc g' hᶜ)
|
||||||
|
with ≈-dec a₂ (f a₂)
|
||||||
|
... | yes _ = a₂≼a
|
||||||
|
... | no _ = stepPreservesLess g' _ _ _ a a≈fa (≼-cong ≈-refl (≈-sym a≈fa) (Monotonicᶠ a₂≼a)) _ _ _
|
||||||
|
|
||||||
|
aᶠ≼ : ∀ (a : A) → a ≈ f a → aᶠ ≼ a
|
||||||
|
aᶠ≼ a a≈fa = stepPreservesLess (suc h) 0 ⊥ᴬ ⊥ᴬ a a≈fa (⊥ᴬ≼ a) (ChainA.done ≈-refl) (+-comm (suc h) 0) (⊥ᴬ≼ (f ⊥ᴬ))
|
||||||
|
|
||||||
|
|
Loading…
Reference in New Issue
Block a user