Library Coq.Logic.SetoidChoice
This module states the functional form of the axiom of choice over
setoids, commonly called extensional axiom of choice [Carlström04],
[Martin-Löf05]. This is obtained by a decomposition of the axiom
into the following components:
Among other results, it entails:
[Carlström04] Jesper Carlström, EM + Ext + AC_int is equivalent to
AC_ext, Mathematical Logic Quaterly, vol 50(3), pp 236-240, 2004.
[Martin-Löf05] Per Martin-Löf, 100 years of Zermelo’s axiom of
choice: what was the problem with it?, lecture notes for KTH/SU
colloquium, 2005.
- classical logic
- relational axiom of choice
- axiom of unique choice
- a limited form of functional extensionality
- proof irrelevance
- choice of a representative in equivalence classes
Require Export ClassicalChoice. Require Export ExtensionalFunctionRepresentative.
Require Import ChoiceFacts.
Require Import ClassicalFacts.
Require Import RelationClasses.
Theorem setoid_choice :
forall A B,
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').
Theorem representative_choice :
forall A (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'.