Library Coq.Logic.ChoiceFacts


Some facts and definitions concerning choice and description in intuitionistic logic.

References:

[Bell] John L. Bell, Choice principles in intuitionistic set theory, unpublished.
[Bell93] John L. Bell, Hilbert's Epsilon Operator in Intuitionistic Type Theories, Mathematical Logic Quarterly, volume 39, 1993.
[Carlström04] Jesper Carlström, EM + Ext + AC_int is equivalent to AC_ext, Mathematical Logic Quaterly, vol 50(3), pp 236-240, 2004.
[Carlström05] Jesper Carlström, Interpreting descriptions in intentional type theory, Journal of Symbolic Logic 70(2):488-514, 2005.
[Werner97] Benjamin Werner, Sets in Types, Types in Sets, TACS, 1997.

Require Import RelationClasses Logic.

Set Implicit Arguments.

Definitions

Choice, reification and description schemes
We make them all polymorphic. Most of them have existentials as conclusion so they require polymorphism otherwise their first application (e.g. to an existential in Set) will fix the level of A.

Section ChoiceSchemes.

Variables A B :Type.

Variable P:A->Prop.

Constructive choice and description

AC_rel = relational form of the (non extensional) axiom of choice (a "set-theoretic" axiom of choice)

Definition RelationalChoice_on :=
  forall R:A->B->Prop,
    (forall x : A, exists y : B, R x y) ->
    (exists R' : A->B->Prop, subrelation R' R /\ forall x, exists! y, R' x y).

AC_fun = functional form of the (non extensional) axiom of choice (a "type-theoretic" axiom of choice)


Definition FunctionalChoice_on_rel (R:A->B->Prop) :=
  (forall x:A, exists y : B, R x y) ->
  exists f : A -> B, (forall x:A, R x (f x)).

Definition FunctionalChoice_on :=
  forall R:A->B->Prop,
    (forall x : A, exists y : B, R x y) ->
    (exists f : A->B, forall x : A, R x (f x)).

AC_fun_dep = functional form of the (non extensional) axiom of choice, with dependent functions
Definition DependentFunctionalChoice_on (A:Type) (B:A -> Type) :=
  forall R:forall x:A, B x -> Prop,
    (forall x:A, exists y : B x, R x y) ->
    (exists f : (forall x:A, B x), forall x:A, R x (f x)).

AC_trunc = axiom of choice for propositional truncations (truncation and quantification commute)
Definition InhabitedForallCommute_on (A : Type) (B : A -> Type) :=
  (forall x, inhabited (B x)) -> inhabited (forall x, B x).

DC_fun = functional form of the dependent axiom of choice

Definition FunctionalDependentChoice_on :=
  forall (R:A->A->Prop),
    (forall x, exists y, R x y) -> forall x0,
    (exists f : nat -> A, f 0 = x0 /\ forall n, R (f n) (f (S n))).

ACw_fun = functional form of the countable axiom of choice

Definition FunctionalCountableChoice_on :=
  forall (R:nat->A->Prop),
    (forall n, exists y, R n y) ->
    (exists f : nat -> A, forall n, R n (f n)).

AC! = functional relation reification (known as axiom of unique choice in topos theory, sometimes called principle of definite description in the context of constructive type theory, sometimes called axiom of no choice)

Definition FunctionalRelReification_on :=
  forall R:A->B->Prop,
    (forall x : A, exists! y : B, R x y) ->
    (exists f : A->B, forall x : A, R x (f x)).

AC_dep! = functional relation reification, with dependent functions see AC!
Definition DependentFunctionalRelReification_on (A:Type) (B:A -> Type) :=
  forall (R:forall x:A, B x -> Prop),
    (forall x:A, exists! y : B x, R x y) ->
    (exists f : (forall x:A, B x), forall x:A, R x (f x)).

AC_fun_repr = functional choice of a representative in an equivalence class


Definition RepresentativeFunctionalChoice_on :=
  forall R:A->A->Prop,
    (Equivalence R) ->
    (exists f : A->A, forall x : A, (R x (f x)) /\ forall x', R x x' -> f x = f x').

AC_fun_setoid = functional form of the (so-called extensional) axiom of choice from setoids

Definition SetoidFunctionalChoice_on :=
  forall R : A -> A -> Prop,
  forall T : A -> B -> Prop,
  Equivalence R ->
  (forall x x' y, R x x' -> T x y -> T x' y) ->
  (forall x, exists y, T x y) ->
  exists f : A -> B, forall x : A, T x (f x) /\ (forall x' : A, R x x' -> f x = f x').

AC_fun_setoid_gen = functional form of the general form of the (so-called extensional) axiom of choice over setoids


Definition GeneralizedSetoidFunctionalChoice_on :=
  forall R : A -> A -> Prop,
  forall S : B -> B -> Prop,
  forall T : A -> B -> Prop,
  Equivalence R ->
  Equivalence S ->
  (forall x x' y y', R x x' -> S y y' -> T x y -> T x' y') ->
  (forall x, exists y, T x y) ->
  exists f : A -> B,
    forall x : A, T x (f x) /\ (forall x' : A, R x x' -> S (f x) (f x')).

AC_fun_setoid_simple = functional form of the (so-called extensional) axiom of choice from setoids on locally compatible relations

Definition SimpleSetoidFunctionalChoice_on A B :=
  forall R : A -> A -> Prop,
  forall T : A -> B -> Prop,
  Equivalence R ->
  (forall x, exists y, forall x', R x x' -> T x' y) ->
  exists f : A -> B, forall x : A, T x (f x) /\ (forall x' : A, R x x' -> f x = f x').

ID_epsilon = constructive version of indefinite description; combined with proof-irrelevance, it may be connected to Carlström's type theory with a constructive indefinite description operator

Definition ConstructiveIndefiniteDescription_on :=
  forall P:A->Prop,
    (exists x, P x) -> { x:A | P x }.

ID_iota = constructive version of definite description; combined with proof-irrelevance, it may be connected to Carlström's and Stenlund's type theory with a constructive definite description operator)

Definition ConstructiveDefiniteDescription_on :=
  forall P:A->Prop,
    (exists! x, P x) -> { x:A | P x }.

Weakly classical choice and description

GAC_rel = guarded relational form of the (non extensional) axiom of choice

Definition GuardedRelationalChoice_on :=
  forall P : A->Prop, forall R : A->B->Prop,
    (forall x : A, P x -> exists y : B, R x y) ->
    (exists R' : A->B->Prop,
      subrelation R' R /\ forall x, P x -> exists! y, R' x y).

GAC_fun = guarded functional form of the (non extensional) axiom of choice

Definition GuardedFunctionalChoice_on :=
  forall P : A->Prop, forall R : A->B->Prop,
    inhabited B ->
    (forall x : A, P x -> exists y : B, R x y) ->
    (exists f : A->B, forall x, P x -> R x (f x)).

GAC! = guarded functional relation reification

Definition GuardedFunctionalRelReification_on :=
  forall P : A->Prop, forall R : A->B->Prop,
    inhabited B ->
    (forall x : A, P x -> exists! y : B, R x y) ->
    (exists f : A->B, forall x : A, P x -> R x (f x)).

OAC_rel = "omniscient" relational form of the (non extensional) axiom of choice

Definition OmniscientRelationalChoice_on :=
  forall R : A->B->Prop,
    exists R' : A->B->Prop,
      subrelation R' R /\ forall x : A, (exists y : B, R x y) -> exists! y, R' x y.

OAC_fun = "omniscient" functional form of the (non extensional) axiom of choice (called AC* in Bell [Bell])

Definition OmniscientFunctionalChoice_on :=
  forall R : A->B->Prop,
    inhabited B ->
    exists f : A->B, forall x : A, (exists y : B, R x y) -> R x (f x).

D_epsilon = (weakly classical) indefinite description principle

Definition EpsilonStatement_on :=
  forall P:A->Prop,
    inhabited A -> { x:A | (exists x, P x) -> P x }.

D_iota = (weakly classical) definite description principle

Definition IotaStatement_on :=
  forall P:A->Prop,
    inhabited A -> { x:A | (exists! x, P x) -> P x }.

End ChoiceSchemes.

Generalized schemes

Notation RelationalChoice :=
  (forall A B : Type, RelationalChoice_on A B).
Notation FunctionalChoice :=
  (forall A B : Type, FunctionalChoice_on A B).
Notation DependentFunctionalChoice :=
  (forall A (B:A->Type), DependentFunctionalChoice_on B).
Notation InhabitedForallCommute :=
  (forall A (B : A -> Type), InhabitedForallCommute_on B).
Notation FunctionalDependentChoice :=
  (forall A : Type, FunctionalDependentChoice_on A).
Notation FunctionalCountableChoice :=
  (forall A : Type, FunctionalCountableChoice_on A).
Notation FunctionalChoiceOnInhabitedSet :=
  (forall A B : Type, inhabited B -> FunctionalChoice_on A B).
Notation FunctionalRelReification :=
  (forall A B : Type, FunctionalRelReification_on A B).
Notation DependentFunctionalRelReification :=
  (forall A (B:A->Type), DependentFunctionalRelReification_on B).
Notation RepresentativeFunctionalChoice :=
  (forall A : Type, RepresentativeFunctionalChoice_on A).
Notation SetoidFunctionalChoice :=
  (forall A B: Type, SetoidFunctionalChoice_on A B).
Notation GeneralizedSetoidFunctionalChoice :=
  (forall A B : Type, GeneralizedSetoidFunctionalChoice_on A B).
Notation SimpleSetoidFunctionalChoice :=
  (forall A B : Type, SimpleSetoidFunctionalChoice_on A B).

Notation GuardedRelationalChoice :=
  (forall A B : Type, GuardedRelationalChoice_on A B).
Notation GuardedFunctionalChoice :=
  (forall A B : Type, GuardedFunctionalChoice_on A B).
Notation GuardedFunctionalRelReification :=
  (forall A B : Type, GuardedFunctionalRelReification_on A B).

Notation OmniscientRelationalChoice :=
  (forall A B : Type, OmniscientRelationalChoice_on A B).
Notation OmniscientFunctionalChoice :=
  (forall A B : Type, OmniscientFunctionalChoice_on A B).

Notation ConstructiveDefiniteDescription :=
  (forall A : Type, ConstructiveDefiniteDescription_on A).
Notation ConstructiveIndefiniteDescription :=
  (forall A : Type, ConstructiveIndefiniteDescription_on A).

Notation IotaStatement :=
  (forall A : Type, IotaStatement_on A).
Notation EpsilonStatement :=
  (forall A : Type, EpsilonStatement_on A).

Subclassical schemes
PI = proof irrelevance
Definition ProofIrrelevance :=
  forall (A:Prop) (a1 a2:A), a1 = a2.

IGP = independence of general premises (an unconstrained generalisation of the constructive principle of independence of premises)
Definition IndependenceOfGeneralPremises :=
  forall (A:Type) (P:A -> Prop) (Q:Prop),
    inhabited A ->
    (Q -> exists x, P x) -> exists x, Q -> P x.

Drinker = drinker's paradox (small form) (called Ex in Bell [Bell])
Definition SmallDrinker'sParadox :=
  forall (A:Type) (P:A -> Prop), inhabited A ->
    exists x, (exists x, P x) -> P x.

EM = excluded-middle
Definition ExcludedMiddle :=
  forall P:Prop, P \/ ~ P.

Extensional schemes
Ext_prop_repr = choice of a representative among extensional propositions
Local Notation ExtensionalPropositionRepresentative :=
  (forall (A:Type),
   exists h : Prop -> Prop,
   forall P : Prop, (P <-> h P) /\ forall Q, (P <-> Q) -> h P = h Q).

Ext_pred_repr = choice of a representative among extensional predicates
Local Notation ExtensionalPredicateRepresentative :=
  (forall (A:Type),
   exists h : (A->Prop) -> (A->Prop),
   forall (P : A -> Prop), (forall x, P x <-> h P x) /\ forall Q, (forall x, P x <-> Q x) -> h P = h Q).

Ext_fun_repr = choice of a representative among extensional functions
Local Notation ExtensionalFunctionRepresentative :=
  (forall (A B:Type),
   exists h : (A->B) -> (A->B),
   forall (f : A -> B), (forall x, f x = h f x) /\ forall g, (forall x, f x = g x) -> h f = h g).

We let also
  • IPL_2 = 2nd-order impredicative minimal predicate logic (with ex. quant.)
  • IPL^2 = 2nd-order functional minimal predicate logic (with ex. quant.)
  • IPL_2^2 = 2nd-order impredicative, 2nd-order functional minimal pred. logic (with ex. quant.)
with no prerequisite on the non-emptiness of domains

Table of contents


1. Definitions
2. IPL_2^2 |- AC_rel + AC! = AC_fun
3.1. typed IPL_2 + Sigma-types + PI |- AC_rel = GAC_rel and IPL_2 |- AC_rel + IGP -> GAC_rel and IPL_2 |- GAC_rel = OAC_rel
3.2. IPL^2 |- AC_fun + IGP = GAC_fun = OAC_fun = AC_fun + Drinker
3.3. D_iota -> ID_iota and D_epsilon <-> ID_epsilon + Drinker
4. Derivability of choice for decidable relations with well-ordered codomain
5. AC_fun = AC_fun_dep = AC_trunc
6. Non contradiction of constructive descriptions wrt functional choices
7. Definite description transports classical logic to the computational world
8. Choice -> Dependent choice -> Countable choice
9.1. AC_fun_setoid = AC_fun + Ext_fun_repr + EM
9.2. AC_fun_setoid = AC_fun + Ext_pred_repr + PI

AC_rel + AC! = AC_fun

We show that the functional formulation of the axiom of Choice (usual formulation in type theory) is equivalent to its relational formulation (only formulation of set theory) + functional relation reification (aka axiom of unique choice, or, principle of (parametric) definite descriptions)
This shows that the axiom of choice can be assumed (under its relational formulation) without known inconsistency with classical logic, though functional relation reification conflicts with classical logic

Connection between the guarded, non guarded and omniscient choices

We show that the guarded formulations of the axiom of choice are equivalent to their "omniscient" variant and comes from the non guarded formulation in presence either of the independence of general premises or subset types (themselves derivable from subtypes thanks to proof- irrelevance)

AC_rel + PI -> GAC_rel and AC_rel + IGP -> GAC_rel and GAC_rel = OAC_rel

OAC_rel = GAC_rel

AC_fun + IGP = GAC_fun = OAC_fun = AC_fun + Drinker

AC_fun + IGP = GAC_fun
AC_fun + Drinker = OAC_fun
This was already observed by Bell [Bell]
OAC_fun = GAC_fun
This is derivable from the intuitionistic equivalence between IGP and Drinker but we give a direct proof

D_iota -> ID_iota and D_epsilon <-> ID_epsilon + Drinker

D_iota -> ID_iota
ID_epsilon + Drinker <-> D_epsilon

Derivability of choice for decidable relations with well-ordered codomain

Countable codomains, such as nat, can be equipped with a well-order, which implies the existence of a least element on inhabited decidable subsets. As a consequence, the relational form of the axiom of choice is derivable on nat for decidable relations.
We show instead that functional relation reification and the functional form of the axiom of choice are equivalent on decidable relations with nat as codomain.

Require Import Wf_nat.
Require Import Decidable.

Lemma classical_denumerable_description_imp_fun_choice :
  forall A:Type,
    FunctionalRelReification_on A nat ->
    forall R:A->nat->Prop,
      (forall x y, decidable (R x y)) -> FunctionalChoice_on_rel R.

AC_fun = AC_fun_dep = AC_trunc

Choice on dependent and non dependent function types are equivalent

The easy part
Deriving choice on product types requires some computation on singleton propositional types, so we need computational conjunction projections and dependent elimination of conjunction and equality

Scheme and_indd := Induction for and Sort Prop.
Scheme eq_indd := Induction for eq Sort Prop.

Definition proj1_inf (A B:Prop) (p : A/\B) :=
  let (a,b) := p in a.

Theorem non_dep_dep_functional_choice :
  FunctionalChoice -> DependentFunctionalChoice.

Functional choice and truncation choice are equivalent

Reification of dependent and non dependent functional relation are equivalent

The easy part
Deriving choice on product types requires some computation on singleton propositional types, so we need computational conjunction projections and dependent elimination of conjunction and equality

Non contradiction of constructive descriptions wrt functional axioms of choice

Non contradiction of indefinite description

Non contradiction of definite description

Remark, the following corollaries morally hold:
Definition In_propositional_context (A:Type) := forall C:Prop, (A -> C) -> C.
Corollary constructive_definite_descr_in_prop_context_iff_fun_reification : In_propositional_context ConstructiveIndefiniteDescription <-> FunctionalChoice.
Corollary constructive_definite_descr_in_prop_context_iff_fun_reification : In_propositional_context ConstructiveDefiniteDescription <-> FunctionalRelReification.
but expecting FunctionalChoice (resp. FunctionalRelReification) to be applied on the same Type universes on both sides of the first (resp. second) equivalence breaks the stratification of universes.

Excluded-middle + definite description => computational excluded-middle

The idea for the following proof comes from [ChicliPottierSimpson02]
Classical logic and axiom of unique choice (i.e. functional relation reification), as shown in [ChicliPottierSimpson02], implies the double-negation of excluded-middle in Set (which is incompatible with the impredicativity of Set).
We adapt the proof to show that constructive definite description transports excluded-middle from Prop to Set.
[ChicliPottierSimpson02] Laurent Chicli, Loïc Pottier, Carlos Simpson, Mathematical Quotients and Quotient Types in Coq, Proceedings of TYPES 2002, Lecture Notes in Computer Science 2646, Springer Verlag.

Require Import Setoid.

Theorem constructive_definite_descr_excluded_middle :
  (forall A : Type, ConstructiveDefiniteDescription_on A) ->
  (forall P:Prop, P \/ ~ P) -> (forall P:Prop, {P} + {~ P}).

Corollary fun_reification_descr_computational_excluded_middle_in_prop_context :
  FunctionalRelReification ->
  (forall P:Prop, P \/ ~ P) ->
  forall C:Prop, ((forall P:Prop, {P} + {~ P}) -> C) -> C.

Choice => Dependent choice => Countable choice

About the axiom of choice over setoids

Consequences of the choice of a representative in an equivalence class

This is a variant of Diaconescu and Goodman-Myhill theorems

AC_fun_setoid = AC_fun_setoid_gen = AC_fun_setoid_simple

AC_fun_setoid = AC! + AC_fun_repr

Note: What characterization to give of RepresentativeFunctionalChoice? A formulation of it as a functional relation would certainly be equivalent to the formulation of SetoidFunctionalChoice as a functional relation, but in their functional forms, SetoidFunctionalChoice seems strictly stronger

AC_fun_setoid = AC_fun + Ext_fun_repr + EM


Import EqNotations.

This is the main theorem in [Carlström04]

Note: all ingredients have a computational meaning when taken in separation. However, to compute with the functional choice, existential quantification has to be thought as a strong existential, which is incompatible with the computational content of excluded-middle

AC_fun_setoid = AC_fun + Ext_pred_repr + PI

Note: all ingredients have a computational meaning when taken in separation. However, to compute with the functional choice, existential quantification has to be thought as a strong existential, which is incompatible with proof-irrelevance which requires existential quantification to be truncated

Compatibility notations