Skip to content

Commit

Permalink
add module foundation.uniformly-decidable-type-families
Browse files Browse the repository at this point in the history
  • Loading branch information
fredrik-bakke committed Jan 7, 2025
1 parent 4f8b9c6 commit efaaa3a
Show file tree
Hide file tree
Showing 6 changed files with 213 additions and 30 deletions.
1 change: 1 addition & 0 deletions src/foundation.lagda.md
Original file line number Diff line number Diff line change
Expand Up @@ -445,6 +445,7 @@ open import foundation.type-arithmetic-standard-pullbacks public
open import foundation.type-arithmetic-unit-type public
open import foundation.type-duality public
open import foundation.type-theoretic-principle-of-choice public
open import foundation.uniformly-decidable-type-families public
open import foundation.unions-subtypes public
open import foundation.uniqueness-image public
open import foundation.uniqueness-quantification public
Expand Down
10 changes: 8 additions & 2 deletions src/foundation/coproduct-types.lagda.md
Original file line number Diff line number Diff line change
Expand Up @@ -176,11 +176,17 @@ module _
{l1 l2 : Level} {A : UU l1} {B : UU l2}
where

noncontractibility-coproduct-is-contr' :
is-contr A is-contr B noncontractibility' (A + B) 1
noncontractibility-coproduct-is-contr' HA HB =
inl (center HA) , inr (center HB) , neq-inl-inr

abstract
is-not-contractible-coproduct-is-contr :
is-contr A is-contr B is-not-contractible (A + B)
is-not-contractible-coproduct-is-contr HA HB HAB =
neq-inl-inr {x = center HA} {y = center HB} (eq-is-contr HAB)
is-not-contractible-coproduct-is-contr HA HB =
is-not-contractible-noncontractibility
( 1 , noncontractibility-coproduct-is-contr' HA HB)
```

### Coproducts of mutually exclusive propositions are propositions
Expand Down
25 changes: 14 additions & 11 deletions src/foundation/decidable-dependent-function-types.lagda.md
Original file line number Diff line number Diff line change
Expand Up @@ -10,6 +10,7 @@ module foundation.decidable-dependent-function-types where
open import foundation.decidable-types
open import foundation.functoriality-dependent-function-types
open import foundation.maybe
open import foundation.uniformly-decidable-type-families
open import foundation.universal-property-coproduct-types
open import foundation.universal-property-maybe
open import foundation.universe-levels
Expand Down Expand Up @@ -37,7 +38,7 @@ We describe conditions under which
is-decidable-Π-uniformly-decidable-family :
{l1 l2 : Level} {A : UU l1} {B : A UU l2}
is-decidable A
(((a : A) B a) + ((a : A) ¬ (B a)))
is-uniformly-decidable-family B
is-decidable ((a : A) (B a))
is-decidable-Π-uniformly-decidable-family (inl a) (inl b) =
inl b
Expand All @@ -52,7 +53,8 @@ is-decidable-Π-uniformly-decidable-family (inr na) _ =
```agda
is-decidable-Π-coproduct :
{l1 l2 l3 : Level} {A : UU l1} {B : UU l2} {C : A + B UU l3}
is-decidable ((a : A) C (inl a)) is-decidable ((b : B) C (inr b))
is-decidable ((a : A) C (inl a))
is-decidable ((b : B) C (inr b))
is-decidable ((x : A + B) C x)
is-decidable-Π-coproduct {C = C} dA dB =
is-decidable-equiv
Expand All @@ -77,15 +79,16 @@ is-decidable-Π-Maybe {B = B} du de =
### Decidability of dependent products over an equivalence

```agda
is-decidable-Π-equiv :
module _
{l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {C : A UU l3} {D : B UU l4}
(e : A ≃ B) (f : (x : A) C x ≃ D (map-equiv e x))
is-decidable ((x : A) C x) is-decidable ((y : B) D y)
is-decidable-Π-equiv {D = D} e f = is-decidable-equiv' (equiv-Π D e f)
(e : A ≃ B) (f : (x : A) C x ≃ D (map-equiv e x))
where

is-decidable-Π-equiv' :
{l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {C : A UU l3} {D : B UU l4}
(e : A ≃ B) (f : (x : A) C x ≃ D (map-equiv e x))
is-decidable ((y : B) D y) is-decidable ((x : A) C x)
is-decidable-Π-equiv' {D = D} e f = is-decidable-equiv (equiv-Π D e f)
is-decidable-Π-equiv :
is-decidable ((x : A) C x) is-decidable ((y : B) D y)
is-decidable-Π-equiv = is-decidable-equiv' (equiv-Π D e f)

is-decidable-Π-equiv' :
is-decidable ((y : B) D y) is-decidable ((x : A) C x)
is-decidable-Π-equiv' = is-decidable-equiv (equiv-Π D e f)
```
30 changes: 14 additions & 16 deletions src/foundation/decidable-dependent-pair-types.lagda.md
Original file line number Diff line number Diff line change
Expand Up @@ -12,6 +12,7 @@ open import foundation.dependent-pair-types
open import foundation.maybe
open import foundation.type-arithmetic-coproduct-types
open import foundation.type-arithmetic-unit-type
open import foundation.uniformly-decidable-type-families
open import foundation.universe-levels

open import foundation-core.coproduct-types
Expand All @@ -36,7 +37,9 @@ We describe conditions under which
```agda
is-decidable-Σ-uniformly-decidable-family :
{l1 l2 : Level} {A : UU l1} {B : A UU l2}
is-decidable A (((a : A) B a) + ((a : A) ¬ B a)) is-decidable (Σ A B)
is-decidable A
is-uniformly-decidable-family B
is-decidable (Σ A B)
is-decidable-Σ-uniformly-decidable-family (inl a) (inl b) =
inl (a , b a)
is-decidable-Σ-uniformly-decidable-family (inl a) (inr b) =
Expand Down Expand Up @@ -65,27 +68,22 @@ is-decidable-Σ-Maybe :
{l1 l2 : Level} {A : UU l1} {B : Maybe A UU l2}
is-decidable (Σ A (B ∘ unit-Maybe)) is-decidable (B exception-Maybe)
is-decidable (Σ (Maybe A) B)
is-decidable-Σ-Maybe {l1} {l2} {A} {B} dA de =
is-decidable-Σ-Maybe {A = A} {B} dA de =
is-decidable-Σ-coproduct B dA
( is-decidable-equiv
( left-unit-law-Σ (B ∘ inr))
( de))
( is-decidable-equiv (left-unit-law-Σ (B ∘ inr)) de)
```

### Decidability of dependent sums over equivalences

```agda
is-decidable-Σ-equiv :
module _
{l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {C : A UU l3} {D : B UU l4}
(e : A ≃ B) (f : (x : A) C x ≃ D (map-equiv e x))
is-decidable (Σ A C) is-decidable (Σ B D)
is-decidable-Σ-equiv {D = D} e f =
is-decidable-equiv' (equiv-Σ D e f)
(e : A ≃ B) (f : (x : A) C x ≃ D (map-equiv e x))
where

is-decidable-Σ-equiv' :
{l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {C : A UU l3} {D : B UU l4}
(e : A ≃ B) (f : (x : A) C x ≃ D (map-equiv e x))
is-decidable (Σ B D) is-decidable (Σ A C)
is-decidable-Σ-equiv' {D = D} e f =
is-decidable-equiv (equiv-Σ D e f)
is-decidable-Σ-equiv : is-decidable (Σ A C) is-decidable (Σ B D)
is-decidable-Σ-equiv = is-decidable-equiv' (equiv-Σ D e f)

is-decidable-Σ-equiv' : is-decidable (Σ B D) is-decidable (Σ A C)
is-decidable-Σ-equiv' = is-decidable-equiv (equiv-Σ D e f)
```
2 changes: 1 addition & 1 deletion src/foundation/path-split-type-families.lagda.md
Original file line number Diff line number Diff line change
Expand Up @@ -40,7 +40,7 @@ This condition is a direct rephrasing of stating that the
[action on identifications](foundation.action-on-identifications-functions.md)
of the first projection map `Σ A P A` has a
[section](foundation-core.sections.md), and in this way is closely related to
the concept of [path-split maps](foundation-core.path-splt-maps.md).
the concept of [path-split maps](foundation-core.path-split-maps.md).

## Definitions

Expand Down
175 changes: 175 additions & 0 deletions src/foundation/uniformly-decidable-type-families.lagda.md
Original file line number Diff line number Diff line change
@@ -0,0 +1,175 @@
# Uniformly decidable type families

```agda
module foundation.uniformly-decidable-type-families where
```

<details><summary>Imports</summary>

```agda
open import foundation.contractible-types
open import foundation.coproduct-types
open import foundation.decidable-types
open import foundation.dependent-pair-types
open import foundation.equality-coproduct-types
open import foundation.inhabited-types
open import foundation.negation
open import foundation.propositional-truncations
open import foundation.propositions
open import foundation.truncated-types
open import foundation.truncation-levels
open import foundation.type-arithmetic-empty-type
open import foundation.universe-levels

open import foundation-core.cartesian-product-types
open import foundation-core.contractible-maps
open import foundation-core.empty-types
open import foundation-core.equivalences
open import foundation-core.function-types
open import foundation-core.homotopies
open import foundation-core.identity-types
```

</details>

## Idea

A type family `B : A 𝒰` is
{{#concept "uniformly decidable" Agda=is-uniformly-decidable-type-family}} if
there either is an element of every fiber `B x`, or every fiber is
[empty](foundation-core.empty-types.md).

## Definitions

### The predicate of being uniformly decidable

```agda
is-uniformly-decidable-family :
{l1 l2 : Level} {A : UU l1} (A UU l2) UU (l1 ⊔ l2)
is-uniformly-decidable-family {A = A} B =
((x : A) B x) + ((x : A) ¬ (B x))
```

## Properties

### The fibers of a uniformly decidable type family are decidable

```agda
is-decidable-is-uniformly-decidable-family :
{l1 l2 : Level} {A : UU l1} {B : A UU l2}
is-uniformly-decidable-family B
(x : A) is-decidable (B x)
is-decidable-is-uniformly-decidable-family (inl f) x = inl (f x)
is-decidable-is-uniformly-decidable-family (inr g) x = inr (g x)
```

### The uniform decidability predicate on a family of truncated types

```agda
module _
{l1 l2 : Level} {A : UU l1} {B : A UU l2}
where

is-contr-is-uniformly-decidable-family-is-inhabited-base'' :
is-contr ((x : A) (B x))
A
is-contr (is-uniformly-decidable-family B)
is-contr-is-uniformly-decidable-family-is-inhabited-base'' H a =
is-contr-equiv
( (x : A) B x)
( right-unit-law-coproduct-is-empty
( (x : A) B x)
( (x : A) ¬ B x)
( λ f f a (center H a)))
( H)

is-contr-is-uniformly-decidable-family-is-inhabited-base' :
((x : A) is-contr (B x))
A
is-contr (is-uniformly-decidable-family B)
is-contr-is-uniformly-decidable-family-is-inhabited-base' H =
is-contr-is-uniformly-decidable-family-is-inhabited-base'' (is-contr-Π H)

is-contr-is-uniformly-decidable-family-is-inhabited-base :
((x : A) is-contr (B x))
is-inhabited A
is-contr (is-uniformly-decidable-family B)
is-contr-is-uniformly-decidable-family-is-inhabited-base H =
rec-trunc-Prop
( is-contr-Prop (is-uniformly-decidable-family B))
( is-contr-is-uniformly-decidable-family-is-inhabited-base' H)
```

### The uniform decidability predicate on a family of propositions

```agda
module _
{l1 l2 : Level} {A : UU l1} {B : A UU l2}
where

is-prop-is-uniformly-decidable-family-is-inhabited-base'' :
is-prop ((x : A) (B x))
A
is-prop (is-uniformly-decidable-family B)
is-prop-is-uniformly-decidable-family-is-inhabited-base'' H a =
is-prop-coproduct
( λ f nf nf a (f a))
( H)
( is-prop-Π (λ x is-prop-neg))

is-prop-is-uniformly-decidable-family-is-inhabited-base' :
((x : A) is-prop (B x))
A
is-prop (is-uniformly-decidable-family B)
is-prop-is-uniformly-decidable-family-is-inhabited-base' H =
is-prop-is-uniformly-decidable-family-is-inhabited-base'' (is-prop-Π H)

is-prop-is-uniformly-decidable-family-is-inhabited-base :
((x : A) is-prop (B x))
is-inhabited A
is-prop (is-uniformly-decidable-family B)
is-prop-is-uniformly-decidable-family-is-inhabited-base H =
rec-trunc-Prop
( is-prop-Prop (is-uniformly-decidable-family B))
( is-prop-is-uniformly-decidable-family-is-inhabited-base' H)
```

### The uniform decidability predicate on a family of truncated types

```agda
module _
{l1 l2 : Level} {A : UU l1} {B : A UU l2}
where

is-trunc-succ-succ-is-uniformly-decidable-family' :
(k : 𝕋)
is-trunc (succ-𝕋 (succ-𝕋 k)) ((x : A) (B x))
is-trunc (succ-𝕋 (succ-𝕋 k)) (is-uniformly-decidable-family B)
is-trunc-succ-succ-is-uniformly-decidable-family' k H =
is-trunc-coproduct k
( H)
( is-trunc-is-prop (succ-𝕋 k) (is-prop-Π (λ x is-prop-neg)))

is-trunc-succ-succ-is-uniformly-decidable-family :
(k : 𝕋)
((x : A) is-trunc (succ-𝕋 (succ-𝕋 k)) (B x))
is-trunc (succ-𝕋 (succ-𝕋 k)) (is-uniformly-decidable-family B)
is-trunc-succ-succ-is-uniformly-decidable-family k H =
is-trunc-succ-succ-is-uniformly-decidable-family' k
( is-trunc-Π (succ-𝕋 (succ-𝕋 k)) H)

is-trunc-is-uniformly-decidable-family-is-inhabited-base :
(k : 𝕋)
((x : A) is-trunc k (B x))
is-inhabited A
is-trunc k (is-uniformly-decidable-family B)
is-trunc-is-uniformly-decidable-family-is-inhabited-base
neg-two-𝕋 =
is-contr-is-uniformly-decidable-family-is-inhabited-base
is-trunc-is-uniformly-decidable-family-is-inhabited-base
( succ-𝕋 neg-two-𝕋) =
is-prop-is-uniformly-decidable-family-is-inhabited-base
is-trunc-is-uniformly-decidable-family-is-inhabited-base
( succ-𝕋 (succ-𝕋 k)) H _ =
is-trunc-succ-succ-is-uniformly-decidable-family k H
```

0 comments on commit efaaa3a

Please sign in to comment.