Skip to content

Commit cb5d2bd

Browse files
authored
[ refactor ] Data.Fin.Properties.decFinSubset (#2793)
* [ refactoring ] decidable `Data.Fin.Properties` * refactor: make `variable`s `P`, `Q` local again * refactor: use flipped contradiction * fix: comment * Restored deleted `import` of `Function.Base.flip`
1 parent 62566af commit cb5d2bd

File tree

1 file changed

+36
-25
lines changed

1 file changed

+36
-25
lines changed

src/Data/Fin/Properties.agda

Lines changed: 36 additions & 25 deletions
Original file line numberDiff line numberDiff line change
@@ -29,7 +29,7 @@ open import Data.Product.Properties using (,-injective)
2929
open import Data.Product.Algebra using (×-cong)
3030
open import Data.Sum.Base as Sum using (_⊎_; inj₁; inj₂; [_,_]; [_,_]′)
3131
open import Data.Sum.Properties using ([,]-map; [,]-∘)
32-
open import Function.Base using (_∘_; id; _$_; flip)
32+
open import Function.Base using (_∘_; id; _$_; flip; const; λ-; _$-)
3333
open import Function.Bundles using (Injection; _↣_; _⇔_; _↔_; mk⇔; mk↔ₛ′)
3434
open import Function.Definitions using (Injective; Surjective)
3535
open import Function.Consequences.Propositional using (contraInjective)
@@ -58,11 +58,12 @@ open import Relation.Unary.Properties using (U?)
5858

5959
private
6060
variable
61-
a : Level
61+
a p q : Level
6262
A : Set a
6363
m n o :
6464
i j : Fin n
6565

66+
6667
------------------------------------------------------------------------
6768
-- Fin
6869
------------------------------------------------------------------------
@@ -996,7 +997,7 @@ pinch-injective {i = suc i} {suc j} {suc k} 1+i≢j 1+i≢k eq =
996997
-- Quantification
997998
------------------------------------------------------------------------
998999

999-
module _ {p} {P : Pred (Fin (suc n)) p} where
1000+
module _ {P : Pred (Fin (suc n)) p} where
10001001

10011002
∀-cons : P zero Π[ P ∘ suc ] Π[ P ]
10021003
∀-cons z s zero = z
@@ -1018,33 +1019,19 @@ module _ {p} {P : Pred (Fin (suc n)) p} where
10181019
⊎⇔∃ : (P zero ⊎ ∃⟨ P ∘ suc ⟩) ⇔ ∃⟨ P ⟩
10191020
⊎⇔∃ = mk⇔ [ ∃-here , ∃-there ] ∃-toSum
10201021

1021-
decFinSubset : {p q} {P : Pred (Fin n) p} {Q : Pred (Fin n) q}
1022-
Decidable Q ( {i} Q i Dec (P i)) Dec (Q ⊆ P)
1023-
decFinSubset {zero} {_} {_} Q? P? = yes λ {}
1024-
decFinSubset {suc n} {P = P} {Q} Q? P?
1025-
with Q? zero | ∀-cons {P = λ x Q x P x}
1026-
... | false because [¬Q0] | cons =
1027-
map′ (λ f {x} cons (⊥-elim ∘ invert [¬Q0]) (λ x f {x}) x)
1028-
(λ f {x} f {suc x})
1029-
(decFinSubset (Q? ∘ suc) P?)
1030-
... | true because [Q0] | cons =
1031-
map′ (uncurry λ P0 rec {x} cons (λ _ P0) (λ x rec {x}) x)
1032-
< _$ invert [Q0] , (λ f {x} f {suc x}) >
1033-
(P? (invert [Q0]) ×-dec decFinSubset (Q? ∘ suc) P?)
1034-
1035-
any? : {p} {P : Pred (Fin n) p} Decidable P Dec (∃ P)
1036-
any? {zero} {P = _} P? = no λ { (() , _) }
1037-
any? {suc n} {P = P} P? = Dec.map ⊎⇔∃ (P? zero ⊎-dec any? (P? ∘ suc))
1038-
1039-
all? : {p} {P : Pred (Fin n) p} Decidable P Dec ( f P f)
1040-
all? P? = map′ (λ ∀p f ∀p tt) (λ ∀p {x} _ ∀p x)
1041-
(decFinSubset U? (λ {f} _ P? f))
1022+
any? : {P : Pred (Fin n) p} Decidable P Dec (∃ P)
1023+
any? {zero} P? = no λ{ (() , _) }
1024+
any? {suc _} P? = Dec.map ⊎⇔∃ (P? zero ⊎-dec any? (P? ∘ suc))
1025+
1026+
all? : {P : Pred (Fin n) p} Decidable P Dec ( i P i)
1027+
all? {zero} P? = yes λ()
1028+
all? {suc _} P? = Dec.map ∀-cons-⇔ (P? zero ×-dec all? (P? ∘ suc))
10421029

10431030
private
10441031
-- A nice computational property of `all?`:
10451032
-- The boolean component of the result is exactly the
10461033
-- obvious fold of boolean tests (`foldr _∧_ true`).
1047-
note : {p} {P : Pred (Fin 3) p} (P? : Decidable P)
1034+
note : {P : Pred (Fin 3) p} (P? : Decidable P)
10481035
λ z Dec.does (all? P?) ≡ z
10491036
note P? = Dec.does (P? 0F) ∧ Dec.does (P? 1F) ∧ Dec.does (P? 2F) ∧ true
10501037
, refl
@@ -1067,6 +1054,30 @@ private
10671054
¬ ( i P i) (∃ λ i ¬ P i)
10681055
¬∀⇒∃¬ n P P? ¬P = map id proj₁ (¬∀⇒∃¬-smallest n P P? ¬P)
10691056

1057+
-- lifting Dec over Unary subset relation
1058+
1059+
decFinSubset : {P : Pred (Fin n) p} {Q : Pred (Fin n) q}
1060+
Decidable Q Q ⊆ Dec ∘ P Dec (Q ⊆ P)
1061+
decFinSubset {zero} {_} {_} Q? P? = yes λ{}
1062+
decFinSubset {suc _} {P = P} {Q = Q} Q? P? = dec[Q⊆P]
1063+
module DecFinSubset where
1064+
Q⊆₀P = Q 0F P 0F
1065+
Q⊆ₛP = Q ∘ suc ⊆ P ∘ suc
1066+
1067+
cons : Q⊆₀P Q⊆ₛP Q ⊆ P
1068+
cons q₀⊆p₀ qₛ⊆pₛ = ∀-cons {P = Q U.⇒ P} q₀⊆p₀ (λ- qₛ⊆pₛ) $-
1069+
1070+
ih : Dec Q⊆ₛP
1071+
ih = decFinSubset (Q? ∘ suc) P?
1072+
1073+
Q⊆P⇒Q⊆ₛP : Q ⊆ P Q⊆ₛP
1074+
Q⊆P⇒Q⊆ₛP q⊆p {x} = q⊆p {suc x}
1075+
1076+
dec[Q⊆P] : Dec (Q ⊆ P)
1077+
dec[Q⊆P] with Q? zero
1078+
... | no ¬q₀ = map′ (cons (contradiction′ ¬q₀)) Q⊆P⇒Q⊆ₛP ih
1079+
... | yes q₀ = map′ (uncurry (cons ∘ const)) < _$ q₀ , Q⊆P⇒Q⊆ₛP > (P? q₀ ×-dec ih)
1080+
10701081
------------------------------------------------------------------------
10711082
-- Properties of functions to and from Fin
10721083
------------------------------------------------------------------------

0 commit comments

Comments
 (0)