From b826881b59d4ff71e51ab215798dbf71e238b025 Mon Sep 17 00:00:00 2001 From: jamesmckinna Date: Sat, 2 Aug 2025 18:01:26 +0100 Subject: [PATCH 1/2] [ add ] clean version of `Data.Fin.Properties.searchMinimalCounterexample` --- CHANGELOG.md | 21 ++++++++++++ src/Data/Fin/Properties.agda | 64 +++++++++++++++++++++++++++++++----- 2 files changed, 76 insertions(+), 9 deletions(-) diff --git a/CHANGELOG.md b/CHANGELOG.md index ef9533efc6..d6b1a555f0 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -70,6 +70,27 @@ Additions to existing modules ≟-≡ : (eq : i ≡ j) → (i ≟ j) ≡ yes eq ≟-≡-refl : (i : Fin n) → (i ≟ i) ≡ yes refl ≟-≢ : (i≢j : i ≢ j) → (i ≟ j) ≡ no i≢j + inject-< : inject j < i + + record MinimalExample (P : Pred (Fin n) p) : Set p where + constructor μ + field + witness : Fin n + example : P witness + minimal : ∀ {j} → .(j < witness) → ¬ P j + + record MinimalCounterexample (P : Pred (Fin n) p) : Set p where + constructor μ + field + witness : Fin n + .contra : ¬ P witness + minimal : ∀ {j} → .(j < witness) → P j + + μ⟨_⟩ : Pred (Fin n) p → Set p + μ⟨¬_⟩ : Pred (Fin n) p → Set p + ¬¬μ⇒μ : Decidable P → μ⟨¬ ∁ P ⟩ → μ⟨ P ⟩ + searchMinimalCounterexample : Decidable P → (∀ i → P i) ⊎ μ⟨¬ P ⟩ + search-μ⟨¬_⟩ : Decidable P → (∀ i → P i) ⊎ μ⟨¬ P ⟩ ``` * In `Data.Nat.Properties`: diff --git a/src/Data/Fin/Properties.agda b/src/Data/Fin/Properties.agda index 48bcf8a5ee..335929a945 100644 --- a/src/Data/Fin/Properties.agda +++ b/src/Data/Fin/Properties.agda @@ -48,16 +48,17 @@ open import Relation.Binary.PropositionalEquality.Properties as ≡ open import Relation.Binary.PropositionalEquality as ≡ using (≡-≟-identity; ≢-≟-identity) open import Relation.Nullary.Decidable as Dec - using (Dec; _because_; yes; no; _×-dec_; _⊎-dec_; map′) + using (Dec; _because_; yes; no; _×-dec_; _⊎-dec_; map′; decidable-stable) open import Relation.Nullary.Negation.Core using (¬_; contradiction) +open import Relation.Nullary.Recomputable using (¬-recompute) open import Relation.Nullary.Reflects using (invert) open import Relation.Unary as U - using (U; Pred; Decidable; _⊆_; Satisfiable; Universal) + using (U; Pred; Decidable; _⊆_; ∁; Satisfiable; Universal) open import Relation.Unary.Properties using (U?) private variable - a : Level + a p q : Level A : Set a m n o : ℕ i j : Fin n @@ -469,6 +470,10 @@ toℕ-inject : ∀ {i : Fin n} (j : Fin′ i) → toℕ (inject j) ≡ toℕ j toℕ-inject {i = suc i} zero = refl toℕ-inject {i = suc i} (suc j) = cong suc (toℕ-inject j) +inject-< : ∀ {i : Fin n} (j : Fin′ i) → inject j < i +inject-< {i = suc i} zero = z Date: Tue, 5 Aug 2025 13:28:15 +0100 Subject: [PATCH 2/2] =?UTF-8?q?use:=20`=CE=A0[=5F]`=20syntax?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- CHANGELOG.md | 4 ++-- src/Data/Fin/Properties.agda | 4 ++-- 2 files changed, 4 insertions(+), 4 deletions(-) diff --git a/CHANGELOG.md b/CHANGELOG.md index d6b1a555f0..b33ff71a91 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -89,8 +89,8 @@ Additions to existing modules μ⟨_⟩ : Pred (Fin n) p → Set p μ⟨¬_⟩ : Pred (Fin n) p → Set p ¬¬μ⇒μ : Decidable P → μ⟨¬ ∁ P ⟩ → μ⟨ P ⟩ - searchMinimalCounterexample : Decidable P → (∀ i → P i) ⊎ μ⟨¬ P ⟩ - search-μ⟨¬_⟩ : Decidable P → (∀ i → P i) ⊎ μ⟨¬ P ⟩ + searchMinimalCounterexample : Decidable P → Π[ P ] ⊎ μ⟨¬ P ⟩ + search-μ⟨¬_⟩ : Decidable P → Π[ P ] ⊎ μ⟨¬ P ⟩ ``` * In `Data.Nat.Properties`: diff --git a/src/Data/Fin/Properties.agda b/src/Data/Fin/Properties.agda index 335929a945..d8bbf76ba1 100644 --- a/src/Data/Fin/Properties.agda +++ b/src/Data/Fin/Properties.agda @@ -1082,7 +1082,7 @@ module _ (P : Pred (Fin n) p) where ¬¬μ⇒μ : ∀ {P : Pred (Fin n) p} → Decidable P → μ⟨¬ ∁ P ⟩ → μ⟨ P ⟩ ¬¬μ⇒μ P? (μ i ¬¬pᵢ ∀[j