diff --git a/CHANGELOG.md b/CHANGELOG.md index b780aa6e1f..3f5cd57927 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -50,6 +50,11 @@ New modules ----------- * `Algebra.Module.Bundles.Raw`: raw bundles for module-like algebraic structures +* Consequences of 'infinite descent' for (accessible elements of) well-founded relations: + ```agda + Induction.InfiniteDescent + ``` + * The unique morphism from the initial, resp. terminal, algebra: ```agda Algebra.Morphism.Construct.Initial @@ -242,6 +247,15 @@ Additions to existing modules * Added new functions in `Data.String.Base`: ```agda map : (Char → Char) → String → String + +* Added new definition in `Relation.Binary.Construct.Closure.Transitive` + ``` + transitive⁻ : Transitive _∼_ → TransClosure _∼_ ⇒ _∼_ + ``` + +* Added new definition in `Relation.Unary` + ``` + Stable : Pred A ℓ → Set _ ``` * In `Function.Bundles`, added `_⟶ₛ_` as a synonym for `Func` that can diff --git a/src/Induction/InfiniteDescent.agda b/src/Induction/InfiniteDescent.agda new file mode 100644 index 0000000000..ed25c880d5 --- /dev/null +++ b/src/Induction/InfiniteDescent.agda @@ -0,0 +1,189 @@ +------------------------------------------------------------------------ +-- The Agda standard library +-- +-- A standard consequence of accessibility/well-foundedness: +-- the impossibility of 'infinite descent' from any (accessible) +-- element x satisfying P to 'smaller' y also satisfying P +------------------------------------------------------------------------ + +{-# OPTIONS --cubical-compatible --safe #-} + +module Induction.InfiniteDescent where + +open import Data.Nat.Base as ℕ using (ℕ; zero; suc) +open import Data.Nat.Properties as ℕ +open import Data.Product.Base using (_,_; proj₁; ∃-syntax; _×_) +open import Function.Base using (_∘_) +open import Induction.WellFounded + using (WellFounded; Acc; acc; acc-inverse; module Some) +open import Level using (Level) +open import Relation.Binary.Core using (Rel) +open import Relation.Binary.Construct.Closure.Transitive +open import Relation.Binary.PropositionalEquality.Core +open import Relation.Nullary.Negation.Core as Negation using (¬_) +open import Relation.Unary + using (Pred; ∁; _∩_; _⊆_; _⇒_; Universal; IUniversal; Stable; Empty) + +private + variable + a r ℓ : Level + A : Set a + f : ℕ → A + _<_ : Rel A r + P : Pred A ℓ + +------------------------------------------------------------------------ +-- Definitions + +InfiniteDescendingSequence : Rel A r → (ℕ → A) → Set _ +InfiniteDescendingSequence _<_ f = ∀ n → f (suc n) < f n + +InfiniteDescendingSequenceFrom : Rel A r → (ℕ → A) → Pred A _ +InfiniteDescendingSequenceFrom _<_ f x = f zero ≡ x × InfiniteDescendingSequence _<_ f + +InfiniteDescendingSequence⁺ : Rel A r → (ℕ → A) → Set _ +InfiniteDescendingSequence⁺ _<_ f = ∀ {m n} → m ℕ.< n → TransClosure _<_ (f n) (f m) + +InfiniteDescendingSequenceFrom⁺ : Rel A r → (ℕ → A) → Pred A _ +InfiniteDescendingSequenceFrom⁺ _<_ f x = f zero ≡ x × InfiniteDescendingSequence⁺ _<_ f + +DescentFrom : Rel A r → Pred A ℓ → Pred A _ +DescentFrom _<_ P x = P x → ∃[ y ] y < x × P y + +Descent : Rel A r → Pred A ℓ → Set _ +Descent _<_ P = ∀ {x} → DescentFrom _<_ P x + +InfiniteDescentFrom : Rel A r → Pred A ℓ → Pred A _ +InfiniteDescentFrom _<_ P x = P x → ∃[ f ] InfiniteDescendingSequenceFrom _<_ f x × ∀ n → P (f n) + +InfiniteDescent : Rel A r → Pred A ℓ → Set _ +InfiniteDescent _<_ P = ∀ {x} → InfiniteDescentFrom _<_ P x + +InfiniteDescentFrom⁺ : Rel A r → Pred A ℓ → Pred A _ +InfiniteDescentFrom⁺ _<_ P x = P x → ∃[ f ] InfiniteDescendingSequenceFrom⁺ _<_ f x × ∀ n → P (f n) + +InfiniteDescent⁺ : Rel A r → Pred A ℓ → Set _ +InfiniteDescent⁺ _<_ P = ∀ {x} → InfiniteDescentFrom⁺ _<_ P x + +NoSmallestCounterExample : Rel A r → Pred A ℓ → Set _ +NoSmallestCounterExample _<_ P = ∀ {x} → Acc _<_ x → DescentFrom (TransClosure _<_) (∁ P) x + +------------------------------------------------------------------------ +-- We can swap between transitively closed and non-transitively closed +-- definitions + +sequence⁺ : InfiniteDescendingSequence (TransClosure _<_) f → + InfiniteDescendingSequence⁺ _<_ f +sequence⁺ {_<_ = _<_} {f = f} seq[f] = seq⁺[f]′ ∘ ℕ.<⇒<′ + where + seq⁺[f]′ : ∀ {m n} → m ℕ.<′ n → TransClosure _<_ (f n) (f m) + seq⁺[f]′ ℕ.<′-base = seq[f] _ + seq⁺[f]′ (ℕ.<′-step m<′n) = seq[f] _ ++ seq⁺[f]′ m<′n + +sequence⁻ : InfiniteDescendingSequence⁺ _<_ f → + InfiniteDescendingSequence (TransClosure _<_) f +sequence⁻ seq[f] = seq[f] ∘ n<1+n + +------------------------------------------------------------------------ +-- Results about unrestricted descent + +module _ (descent : Descent _<_ P) where + + descent∧acc⇒infiniteDescentFrom : (Acc _<_) ⊆ (InfiniteDescentFrom _<_ P) + descent∧acc⇒infiniteDescentFrom {x} = + Some.wfRec (InfiniteDescentFrom _<_ P) rec x + where + rec : _ + rec y rec[y] py + with z , z_) open import Level using (Level; _⊔_; 0ℓ; suc; Lift) -open import Relation.Nullary as Nullary using (¬_; Dec; True) +open import Relation.Nullary.Decidable.Core using (Dec; True) +open import Relation.Nullary as Nullary using (¬_) open import Relation.Binary.PropositionalEquality.Core using (_≡_) private @@ -173,7 +174,7 @@ Irrelevant P = ∀ {x} → Nullary.Irrelevant (P x) Recomputable : Pred A ℓ → Set _ Recomputable P = ∀ {x} → Nullary.Recomputable (P x) --- Weak Decidability +-- Stability - instances of P are stable wrt double negation Stable : Pred A ℓ → Set _ Stable P = ∀ x → Nullary.Stable (P x)