Built with Alectryon, running Coq+SerAPI v8.10.0+0.7.0. Coq sources are in this panel; goals and messages will appear in the other. Bubbles () indicate interactive fragments: hover for details, tap to reveal contents. Use Ctrl+↑ Ctrl+↓ to navigate, Ctrl+🖱️ to focus.
(************************************************************************)
(*         *   The Coq Proof Assistant / The Coq Development Team       *)
(*  v      *   INRIA, CNRS and contributors - Copyright 1999-2018       *)
(* <O___,, *       (see CREDITS file for the list of authors)           *)
(*   \VV/  **************************************************************)
(*    //   *    This file is distributed under the terms of the         *)
(*         *     GNU Lesser General Public License Version 2.1          *)
(*         *     (see LICENSE file for the text of the license)         *)
(************************************************************************)
(************************************************************************)
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.
Local Unset Intuition Negation Unfolding.

(**********************************************************************)

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.
(* Set Universe Polymorphism. *)

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)
(* Note: This is called Type-Theoretic Description Axiom (TTDA) in
   [[Werner97]] (using a non-standard meaning of "description"). This
   is called intensional axiom of choice (AC_int) in [[Carlström04]] *)

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
(* Note: This is called Type-Theoretic Choice Axiom (TTCA) in
   [[Werner97]] (by reference to the extensional set-theoretic
   formulation of choice); Note also a typo in its intended
   formulation in [[Werner97]]. *)

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
(* Note: This is called extensional axiom of choice (AC_ext) in
   [[Carlström04]]. *)

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
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
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
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
with no prerequisite on the non-emptiness of domains
(**********************************************************************)

Table of contents

(* This is very fragile. *)
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

forall A B : Type, FunctionalRelReification_on A B -> RelationalChoice_on A B -> FunctionalChoice_on A B

forall A B : Type, FunctionalRelReification_on A B -> RelationalChoice_on A B -> FunctionalChoice_on A B
A, B:Type
Descr:FunctionalRelReification_on A B
RelCh:RelationalChoice_on A B
R:A -> B -> Prop
H:forall x : A, exists y : B, R x y

exists f : A -> B, forall x : A, R x (f x)
A, B:Type
Descr:FunctionalRelReification_on A B
RelCh:RelationalChoice_on A B
R:A -> B -> Prop
H:forall x : A, exists y : B, R x y
R':A -> B -> Prop
HR'R:subrelation R' R
H0:forall x : A, exists ! y : B, R' x y

exists f : A -> B, forall x : A, R x (f x)
A, B:Type
Descr:FunctionalRelReification_on A B
RelCh:RelationalChoice_on A B
R:A -> B -> Prop
H:forall x : A, exists y : B, R x y
R':A -> B -> Prop
HR'R:subrelation R' R
H0:forall x : A, exists ! y : B, R' x y

forall x : A, exists ! y : B, R' x y
A, B:Type
Descr:FunctionalRelReification_on A B
RelCh:RelationalChoice_on A B
R:A -> B -> Prop
H:forall x : A, exists y : B, R x y
R':A -> B -> Prop
HR'R:subrelation R' R
H0:forall x : A, exists ! y : B, R' x y
f:A -> B
Hf:forall x : A, R' x (f x)
exists f0 : A -> B, forall x : A, R x (f0 x)
A, B:Type
Descr:FunctionalRelReification_on A B
RelCh:RelationalChoice_on A B
R:A -> B -> Prop
H:forall x : A, exists y : B, R x y
R':A -> B -> Prop
HR'R:subrelation R' R
H0:forall x : A, exists ! y : B, R' x y
f:A -> B
Hf:forall x : A, R' x (f x)

exists f0 : A -> B, forall x : A, R x (f0 x)
A, B:Type
Descr:FunctionalRelReification_on A B
RelCh:RelationalChoice_on A B
R:A -> B -> Prop
H:forall x0 : A, exists y : B, R x0 y
R':A -> B -> Prop
HR'R:subrelation R' R
H0:forall x0 : A, exists ! y : B, R' x0 y
f:A -> B
Hf:forall x0 : A, R' x0 (f x0)
x:A

R x (f x)
A, B:Type
Descr:FunctionalRelReification_on A B
RelCh:RelationalChoice_on A B
R:A -> B -> Prop
H:forall x0 : A, exists y0 : B, R x0 y0
R':A -> B -> Prop
HR'R:subrelation R' R
H0:forall x0 : A, exists ! y0 : B, R' x0 y0
f:A -> B
Hf:forall x0 : A, R' x0 (f x0)
x:A
y:B
HR'xy:R' x y
Huniq:forall x' : B, R' x x' -> y = x'

R x (f x)
A, B:Type
Descr:FunctionalRelReification_on A B
RelCh:RelationalChoice_on A B
R:A -> B -> Prop
H:forall x0 : A, exists y0 : B, R x0 y0
R':A -> B -> Prop
HR'R:subrelation R' R
H0:forall x0 : A, exists ! y0 : B, R' x0 y0
f:A -> B
Hf:forall x0 : A, R' x0 (f x0)
x:A
y:B
HR'xy:R' x y
Huniq:forall x' : B, R' x x' -> y = x'

R x y
apply HR'R; assumption. Qed.

forall A B : Type, FunctionalChoice_on A B -> RelationalChoice_on A B

forall A B : Type, FunctionalChoice_on A B -> RelationalChoice_on A B
A, B:Type
FunCh:FunctionalChoice_on A B
R:A -> B -> Prop
H:forall x : A, exists y : B, R x y

exists R' : A -> B -> Prop, subrelation R' R /\ (forall x : A, exists ! y : B, R' x y)
A, B:Type
FunCh:FunctionalChoice_on A B
R:A -> B -> Prop
H:forall x : A, exists y : B, R x y
f:A -> B
H0:forall x : A, R x (f x)

exists R' : A -> B -> Prop, subrelation R' R /\ (forall x : A, exists ! y : B, R' x y)
A, B:Type
FunCh:FunctionalChoice_on A B
R:A -> B -> Prop
H:forall x : A, exists y : B, R x y
f:A -> B
H0:forall x : A, R x (f x)

subrelation (fun (x : A) (y : B) => f x = y) R /\ (forall x : A, exists ! y : B, f x = y)
A, B:Type
FunCh:FunctionalChoice_on A B
R:A -> B -> Prop
H:forall x : A, exists y : B, R x y
f:A -> B
H0:forall x : A, R x (f x)

subrelation (fun (x : A) (y : B) => f x = y) R
A, B:Type
FunCh:FunctionalChoice_on A B
R:A -> B -> Prop
H:forall x : A, exists y : B, R x y
f:A -> B
H0:forall x : A, R x (f x)
forall x : A, exists ! y : B, f x = y
A, B:Type
FunCh:FunctionalChoice_on A B
R:A -> B -> Prop
H:forall x : A, exists y : B, R x y
f:A -> B
H0:forall x : A, R x (f x)

forall x : A, exists ! y : B, f x = y
A, B:Type
FunCh:FunctionalChoice_on A B
R:A -> B -> Prop
H:forall x0 : A, exists y : B, R x0 y
f:A -> B
H0:forall x0 : A, R x0 (f x0)
x:A

f x = f x
A, B:Type
FunCh:FunctionalChoice_on A B
R:A -> B -> Prop
H:forall x0 : A, exists y : B, R x0 y
f:A -> B
H0:forall x0 : A, R x0 (f x0)
x:A
forall x' : B, f x = x' -> f x = x'
A, B:Type
FunCh:FunctionalChoice_on A B
R:A -> B -> Prop
H:forall x0 : A, exists y : B, R x0 y
f:A -> B
H0:forall x0 : A, R x0 (f x0)
x:A

forall x' : B, f x = x' -> f x = x'
trivial. Qed.

forall A B : Type, FunctionalChoice_on A B -> FunctionalRelReification_on A B

forall A B : Type, FunctionalChoice_on A B -> FunctionalRelReification_on A B
A, B:Type
FunCh:FunctionalChoice_on A B
R:A -> B -> Prop
H:forall x : A, exists ! y : B, R x y

exists f : A -> B, forall x : A, R x (f x)
A, B:Type
FunCh:FunctionalChoice_on A B
R:A -> B -> Prop
H:forall x : A, exists ! y : B, R x y

forall x : A, exists y : B, R x y
A, B:Type
FunCh:FunctionalChoice_on A B
R:A -> B -> Prop
H:forall x : A, exists ! y : B, R x y
f:A -> B
H0:forall x : A, R x (f x)
exists f0 : A -> B, forall x : A, R x (f0 x)
(* 1 *)
A, B:Type
FunCh:FunctionalChoice_on A B
R:A -> B -> Prop
H:forall x0 : A, exists ! y : B, R x0 y
x:A

exists y : B, R x y
A, B:Type
FunCh:FunctionalChoice_on A B
R:A -> B -> Prop
H:forall x : A, exists ! y : B, R x y
f:A -> B
H0:forall x : A, R x (f x)
exists f0 : A -> B, forall x : A, R x (f0 x)
A, B:Type
FunCh:FunctionalChoice_on A B
R:A -> B -> Prop
H:forall x0 : A, exists ! y0 : B, R x0 y0
x:A
y:B
HRxy:R x y

exists y0 : B, R x y0
A, B:Type
FunCh:FunctionalChoice_on A B
R:A -> B -> Prop
H:forall x : A, exists ! y : B, R x y
f:A -> B
H0:forall x : A, R x (f x)
exists f0 : A -> B, forall x : A, R x (f0 x)
A, B:Type
FunCh:FunctionalChoice_on A B
R:A -> B -> Prop
H:forall x : A, exists ! y : B, R x y
f:A -> B
H0:forall x : A, R x (f x)

exists f0 : A -> B, forall x : A, R x (f0 x)
(* 2 *) exists f; exact H0. Qed.

forall A B : Type, FunctionalChoice_on A B <-> RelationalChoice_on A B /\ FunctionalRelReification_on A B

forall A B : Type, FunctionalChoice_on A B <-> RelationalChoice_on A B /\ FunctionalRelReification_on A B
A, B:Type

FunctionalChoice_on A B <-> RelationalChoice_on A B /\ FunctionalRelReification_on A B
A, B:Type

FunctionalChoice_on A B -> RelationalChoice_on A B /\ FunctionalRelReification_on A B
A, B:Type
RelationalChoice_on A B /\ FunctionalRelReification_on A B -> FunctionalChoice_on A B
A, B:Type

RelationalChoice_on A B /\ FunctionalRelReification_on A B -> FunctionalChoice_on A B
intros [H H0]; exact (functional_rel_reification_and_rel_choice_imp_fun_choice H0 H). Qed. (**********************************************************************)

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


RelationalChoice -> ProofIrrelevance -> GuardedRelationalChoice

RelationalChoice -> ProofIrrelevance -> GuardedRelationalChoice
rel_choice:RelationalChoice
proof_irrel:ProofIrrelevance

GuardedRelationalChoice
rel_choice:RelationalChoice
proof_irrel:ProofIrrelevance
A, B:Type
P:A -> Prop
R:A -> B -> Prop
H:forall x : A, P x -> exists y : B, R x y

exists R' : A -> B -> Prop, subrelation R' R /\ (forall x : A, P x -> exists ! y : B, R' x y)
rel_choice:RelationalChoice
proof_irrel:ProofIrrelevance
A, B:Type
P:A -> Prop
R:A -> B -> Prop
H:forall x : A, P x -> exists y : B, R x y

forall x : {x : A & P x}, exists y : B, R (projT1 x) y
rel_choice:RelationalChoice
proof_irrel:ProofIrrelevance
A, B:Type
P:A -> Prop
R:A -> B -> Prop
H:forall x : A, P x -> exists y : B, R x y
R':{x : A & P x} -> B -> Prop
HR'R:subrelation R' (fun (x : {x : A & P x}) (y : B) => R (projT1 x) y)
H0:forall x : {x : A & P x}, exists ! y : B, R' x y
exists R'0 : A -> B -> Prop, subrelation R'0 R /\ (forall x : A, P x -> exists ! y : B, R'0 x y)
rel_choice:RelationalChoice
proof_irrel:ProofIrrelevance
A, B:Type
P:A -> Prop
R:A -> B -> Prop
H:forall x0 : A, P x0 -> exists y : B, R x0 y
x:A
HPx:P x

exists y : B, R (projT1 (existT P x HPx)) y
rel_choice:RelationalChoice
proof_irrel:ProofIrrelevance
A, B:Type
P:A -> Prop
R:A -> B -> Prop
H:forall x : A, P x -> exists y : B, R x y
R':{x : A & P x} -> B -> Prop
HR'R:subrelation R' (fun (x : {x : A & P x}) (y : B) => R (projT1 x) y)
H0:forall x : {x : A & P x}, exists ! y : B, R' x y
exists R'0 : A -> B -> Prop, subrelation R'0 R /\ (forall x : A, P x -> exists ! y : B, R'0 x y)
rel_choice:RelationalChoice
proof_irrel:ProofIrrelevance
A, B:Type
P:A -> Prop
R:A -> B -> Prop
H:forall x0 : A, P x0 -> exists y0 : B, R x0 y0
x:A
HPx:P x
y:B
HRxy:R x y

exists y0 : B, R (projT1 (existT P x HPx)) y0
rel_choice:RelationalChoice
proof_irrel:ProofIrrelevance
A, B:Type
P:A -> Prop
R:A -> B -> Prop
H:forall x : A, P x -> exists y : B, R x y
R':{x : A & P x} -> B -> Prop
HR'R:subrelation R' (fun (x : {x : A & P x}) (y : B) => R (projT1 x) y)
H0:forall x : {x : A & P x}, exists ! y : B, R' x y
exists R'0 : A -> B -> Prop, subrelation R'0 R /\ (forall x : A, P x -> exists ! y : B, R'0 x y)
rel_choice:RelationalChoice
proof_irrel:ProofIrrelevance
A, B:Type
P:A -> Prop
R:A -> B -> Prop
H:forall x : A, P x -> exists y : B, R x y
R':{x : A & P x} -> B -> Prop
HR'R:subrelation R' (fun (x : {x : A & P x}) (y : B) => R (projT1 x) y)
H0:forall x : {x : A & P x}, exists ! y : B, R' x y

exists R'0 : A -> B -> Prop, subrelation R'0 R /\ (forall x : A, P x -> exists ! y : B, R'0 x y)
rel_choice:RelationalChoice
proof_irrel:ProofIrrelevance
A, B:Type
P:A -> Prop
R:A -> B -> Prop
H:forall x : A, P x -> exists y : B, R x y
R':{x : A & P x} -> B -> Prop
HR'R:subrelation R' (fun (x : {x : A & P x}) (y : B) => R (projT1 x) y)
H0:forall x : {x : A & P x}, exists ! y : B, R' x y
R'':=fun (x : A) (y : B) => exists H1 : P x, R' (existT P x H1) y:A -> B -> Prop

exists R'0 : A -> B -> Prop, subrelation R'0 R /\ (forall x : A, P x -> exists ! y : B, R'0 x y)
rel_choice:RelationalChoice
proof_irrel:ProofIrrelevance
A, B:Type
P:A -> Prop
R:A -> B -> Prop
H:forall x : A, P x -> exists y : B, R x y
R':{x : A & P x} -> B -> Prop
HR'R:subrelation R' (fun (x : {x : A & P x}) (y : B) => R (projT1 x) y)
H0:forall x : {x : A & P x}, exists ! y : B, R' x y
R'':=fun (x : A) (y : B) => exists H1 : P x, R' (existT P x H1) y:A -> B -> Prop

subrelation R'' R
rel_choice:RelationalChoice
proof_irrel:ProofIrrelevance
A, B:Type
P:A -> Prop
R:A -> B -> Prop
H:forall x : A, P x -> exists y : B, R x y
R':{x : A & P x} -> B -> Prop
HR'R:subrelation R' (fun (x : {x : A & P x}) (y : B) => R (projT1 x) y)
H0:forall x : {x : A & P x}, exists ! y : B, R' x y
R'':=fun (x : A) (y : B) => exists H1 : P x, R' (existT P x H1) y:A -> B -> Prop
forall x : A, P x -> exists ! y : B, R'' x y
rel_choice:RelationalChoice
proof_irrel:ProofIrrelevance
A, B:Type
P:A -> Prop
R:A -> B -> Prop
H:forall x0 : A, P x0 -> exists y0 : B, R x0 y0
R':{x : A & P x} -> B -> Prop
HR'R:subrelation R' (fun (x0 : {x : A & P x}) (y0 : B) => R (projT1 x0) y0)
H0:forall x0 : {x : A & P x}, exists ! y0 : B, R' x0 y0
R'':=fun (x0 : A) (y0 : B) => exists H1 : P x0, R' (existT P x0 H1) y0:A -> B -> Prop
x:A
y:B
HPx:P x
HR'xy:R' (existT P x HPx) y

R x y
rel_choice:RelationalChoice
proof_irrel:ProofIrrelevance
A, B:Type
P:A -> Prop
R:A -> B -> Prop
H:forall x : A, P x -> exists y : B, R x y
R':{x : A & P x} -> B -> Prop
HR'R:subrelation R' (fun (x : {x : A & P x}) (y : B) => R (projT1 x) y)
H0:forall x : {x : A & P x}, exists ! y : B, R' x y
R'':=fun (x : A) (y : B) => exists H1 : P x, R' (existT P x H1) y:A -> B -> Prop
forall x : A, P x -> exists ! y : B, R'' x y
rel_choice:RelationalChoice
proof_irrel:ProofIrrelevance
A, B:Type
P:A -> Prop
R:A -> B -> Prop
H:forall x : A, P x -> exists y : B, R x y
R':{x : A & P x} -> B -> Prop
HR'R:subrelation R' (fun (x : {x : A & P x}) (y : B) => R (projT1 x) y)
H0:forall x : {x : A & P x}, exists ! y : B, R' x y
R'':=fun (x : A) (y : B) => exists H1 : P x, R' (existT P x H1) y:A -> B -> Prop

forall x : A, P x -> exists ! y : B, R'' x y
rel_choice:RelationalChoice
proof_irrel:ProofIrrelevance
A, B:Type
P:A -> Prop
R:A -> B -> Prop
H:forall x0 : A, P x0 -> exists y : B, R x0 y
R':{x : A & P x} -> B -> Prop
HR'R:subrelation R' (fun (x0 : {x : A & P x}) (y : B) => R (projT1 x0) y)
H0:forall x0 : {x : A & P x}, exists ! y : B, R' x0 y
R'':=fun (x0 : A) (y : B) => exists H1 : P x0, R' (existT P x0 H1) y:A -> B -> Prop
x:A
HPx:P x

exists ! y : B, R'' x y
rel_choice:RelationalChoice
proof_irrel:ProofIrrelevance
A, B:Type
P:A -> Prop
R:A -> B -> Prop
H:forall x0 : A, P x0 -> exists y0 : B, R x0 y0
R':{x : A & P x} -> B -> Prop
HR'R:subrelation R' (fun (x0 : {x : A & P x}) (y0 : B) => R (projT1 x0) y0)
H0:forall x0 : {x : A & P x}, exists ! y0 : B, R' x0 y0
R'':=fun (x0 : A) (y0 : B) => exists H1 : P x0, R' (existT P x0 H1) y0:A -> B -> Prop
x:A
HPx:P x
y:B
HR'xy:R' (existT P x HPx) y
Huniq:forall x' : B, R' (existT P x HPx) x' -> y = x'

exists ! y0 : B, R'' x y0
rel_choice:RelationalChoice
proof_irrel:ProofIrrelevance
A, B:Type
P:A -> Prop
R:A -> B -> Prop
H:forall x0 : A, P x0 -> exists y0 : B, R x0 y0
R':{x : A & P x} -> B -> Prop
HR'R:subrelation R' (fun (x0 : {x : A & P x}) (y0 : B) => R (projT1 x0) y0)
H0:forall x0 : {x : A & P x}, exists ! y0 : B, R' x0 y0
R'':=fun (x0 : A) (y0 : B) => exists H1 : P x0, R' (existT P x0 H1) y0:A -> B -> Prop
x:A
HPx:P x
y:B
HR'xy:R' (existT P x HPx) y
Huniq:forall x' : B, R' (existT P x HPx) x' -> y = x'

R'' x y
rel_choice:RelationalChoice
proof_irrel:ProofIrrelevance
A, B:Type
P:A -> Prop
R:A -> B -> Prop
H:forall x0 : A, P x0 -> exists y0 : B, R x0 y0
R':{x : A & P x} -> B -> Prop
HR'R:subrelation R' (fun (x0 : {x : A & P x}) (y0 : B) => R (projT1 x0) y0)
H0:forall x0 : {x : A & P x}, exists ! y0 : B, R' x0 y0
R'':=fun (x0 : A) (y0 : B) => exists H1 : P x0, R' (existT P x0 H1) y0:A -> B -> Prop
x:A
HPx:P x
y:B
HR'xy:R' (existT P x HPx) y
Huniq:forall x' : B, R' (existT P x HPx) x' -> y = x'
forall x' : B, R'' x x' -> y = x'
rel_choice:RelationalChoice
proof_irrel:ProofIrrelevance
A, B:Type
P:A -> Prop
R:A -> B -> Prop
H:forall x0 : A, P x0 -> exists y0 : B, R x0 y0
R':{x : A & P x} -> B -> Prop
HR'R:subrelation R' (fun (x0 : {x : A & P x}) (y0 : B) => R (projT1 x0) y0)
H0:forall x0 : {x : A & P x}, exists ! y0 : B, R' x0 y0
R'':=fun (x0 : A) (y0 : B) => exists H1 : P x0, R' (existT P x0 H1) y0:A -> B -> Prop
x:A
HPx:P x
y:B
HR'xy:R' (existT P x HPx) y
Huniq:forall x' : B, R' (existT P x HPx) x' -> y = x'

forall x' : B, R'' x x' -> y = x'
rel_choice:RelationalChoice
proof_irrel:ProofIrrelevance
A, B:Type
P:A -> Prop
R:A -> B -> Prop
H:forall x0 : A, P x0 -> exists y0 : B, R x0 y0
R':{x : A & P x} -> B -> Prop
HR'R:subrelation R' (fun (x0 : {x : A & P x}) (y0 : B) => R (projT1 x0) y0)
H0:forall x0 : {x : A & P x}, exists ! y0 : B, R' x0 y0
R'':=fun (x0 : A) (y0 : B) => exists H1 : P x0, R' (existT P x0 H1) y0:A -> B -> Prop
x:A
HPx:P x
y:B
HR'xy:R' (existT P x HPx) y
Huniq:forall x' : B, R' (existT P x HPx) x' -> y = x'
y':B
H'Px:P x
HR'xy':R' (existT P x H'Px) y'

y = y'
rel_choice:RelationalChoice
proof_irrel:ProofIrrelevance
A, B:Type
P:A -> Prop
R:A -> B -> Prop
H:forall x0 : A, P x0 -> exists y0 : B, R x0 y0
R':{x : A & P x} -> B -> Prop
HR'R:subrelation R' (fun (x0 : {x : A & P x}) (y0 : B) => R (projT1 x0) y0)
H0:forall x0 : {x : A & P x}, exists ! y0 : B, R' x0 y0
R'':=fun (x0 : A) (y0 : B) => exists H1 : P x0, R' (existT P x0 H1) y0:A -> B -> Prop
x:A
HPx:P x
y:B
HR'xy:R' (existT P x HPx) y
Huniq:forall x' : B, R' (existT P x HPx) x' -> y = x'
y':B
H'Px:P x
HR'xy':R' (existT P x H'Px) y'

R' (existT P x HPx) y'
rewrite proof_irrel with (a1 := HPx) (a2 := H'Px); exact HR'xy'. Qed.

forall A B : Type, inhabited B -> RelationalChoice_on A B -> IndependenceOfGeneralPremises -> GuardedRelationalChoice_on A B

forall A B : Type, inhabited B -> RelationalChoice_on A B -> IndependenceOfGeneralPremises -> GuardedRelationalChoice_on A B
A, B:Type
Inh:inhabited B
AC_rel:RelationalChoice_on A B
IndPrem:IndependenceOfGeneralPremises
P:A -> Prop
R:A -> B -> Prop
H:forall x : A, P x -> exists y : B, R x y

exists R' : A -> B -> Prop, subrelation R' R /\ (forall x : A, P x -> exists ! y : B, R' x y)
A, B:Type
Inh:inhabited B
AC_rel:RelationalChoice_on A B
IndPrem:IndependenceOfGeneralPremises
P:A -> Prop
R:A -> B -> Prop
H:forall x : A, P x -> exists y : B, R x y

forall x : A, exists y : B, P x -> R x y
A, B:Type
Inh:inhabited B
AC_rel:RelationalChoice_on A B
IndPrem:IndependenceOfGeneralPremises
P:A -> Prop
R:A -> B -> Prop
H:forall x : A, P x -> exists y : B, R x y
R':A -> B -> Prop
HR'R:subrelation R' (fun (x : A) (y : B) => P x -> R x y)
H0:forall x : A, exists ! y : B, R' x y
exists R'0 : A -> B -> Prop, subrelation R'0 R /\ (forall x : A, P x -> exists ! y : B, R'0 x y)
A, B:Type
Inh:inhabited B
AC_rel:RelationalChoice_on A B
IndPrem:IndependenceOfGeneralPremises
P:A -> Prop
R:A -> B -> Prop
H:forall x0 : A, P x0 -> exists y : B, R x0 y
x:A

exists y : B, P x -> R x y
A, B:Type
Inh:inhabited B
AC_rel:RelationalChoice_on A B
IndPrem:IndependenceOfGeneralPremises
P:A -> Prop
R:A -> B -> Prop
H:forall x : A, P x -> exists y : B, R x y
R':A -> B -> Prop
HR'R:subrelation R' (fun (x : A) (y : B) => P x -> R x y)
H0:forall x : A, exists ! y : B, R' x y
exists R'0 : A -> B -> Prop, subrelation R'0 R /\ (forall x : A, P x -> exists ! y : B, R'0 x y)
A, B:Type
Inh:inhabited B
AC_rel:RelationalChoice_on A B
IndPrem:IndependenceOfGeneralPremises
P:A -> Prop
R:A -> B -> Prop
H:forall x0 : A, P x0 -> exists y : B, R x0 y
x:A

inhabited B
A, B:Type
Inh:inhabited B
AC_rel:RelationalChoice_on A B
IndPrem:IndependenceOfGeneralPremises
P:A -> Prop
R:A -> B -> Prop
H:forall x0 : A, P x0 -> exists y : B, R x0 y
x:A
P x -> exists x0 : B, R x x0
A, B:Type
Inh:inhabited B
AC_rel:RelationalChoice_on A B
IndPrem:IndependenceOfGeneralPremises
P:A -> Prop
R:A -> B -> Prop
H:forall x : A, P x -> exists y : B, R x y
R':A -> B -> Prop
HR'R:subrelation R' (fun (x : A) (y : B) => P x -> R x y)
H0:forall x : A, exists ! y : B, R' x y
exists R'0 : A -> B -> Prop, subrelation R'0 R /\ (forall x : A, P x -> exists ! y : B, R'0 x y)
A, B:Type
Inh:inhabited B
AC_rel:RelationalChoice_on A B
IndPrem:IndependenceOfGeneralPremises
P:A -> Prop
R:A -> B -> Prop
H:forall x0 : A, P x0 -> exists y : B, R x0 y
x:A

P x -> exists x0 : B, R x x0
A, B:Type
Inh:inhabited B
AC_rel:RelationalChoice_on A B
IndPrem:IndependenceOfGeneralPremises
P:A -> Prop
R:A -> B -> Prop
H:forall x : A, P x -> exists y : B, R x y
R':A -> B -> Prop
HR'R:subrelation R' (fun (x : A) (y : B) => P x -> R x y)
H0:forall x : A, exists ! y : B, R' x y
exists R'0 : A -> B -> Prop, subrelation R'0 R /\ (forall x : A, P x -> exists ! y : B, R'0 x y)
A, B:Type
Inh:inhabited B
AC_rel:RelationalChoice_on A B
IndPrem:IndependenceOfGeneralPremises
P:A -> Prop
R:A -> B -> Prop
H:forall x0 : A, P x0 -> exists y : B, R x0 y
x:A
Hx:P x

exists x0 : B, R x x0
A, B:Type
Inh:inhabited B
AC_rel:RelationalChoice_on A B
IndPrem:IndependenceOfGeneralPremises
P:A -> Prop
R:A -> B -> Prop
H:forall x : A, P x -> exists y : B, R x y
R':A -> B -> Prop
HR'R:subrelation R' (fun (x : A) (y : B) => P x -> R x y)
H0:forall x : A, exists ! y : B, R' x y
exists R'0 : A -> B -> Prop, subrelation R'0 R /\ (forall x : A, P x -> exists ! y : B, R'0 x y)
A, B:Type
Inh:inhabited B
AC_rel:RelationalChoice_on A B
IndPrem:IndependenceOfGeneralPremises
P:A -> Prop
R:A -> B -> Prop
H:forall x : A, P x -> exists y : B, R x y
R':A -> B -> Prop
HR'R:subrelation R' (fun (x : A) (y : B) => P x -> R x y)
H0:forall x : A, exists ! y : B, R' x y

exists R'0 : A -> B -> Prop, subrelation R'0 R /\ (forall x : A, P x -> exists ! y : B, R'0 x y)
A, B:Type
Inh:inhabited B
AC_rel:RelationalChoice_on A B
IndPrem:IndependenceOfGeneralPremises
P:A -> Prop
R:A -> B -> Prop
H:forall x : A, P x -> exists y : B, R x y
R':A -> B -> Prop
HR'R:subrelation R' (fun (x : A) (y : B) => P x -> R x y)
H0:forall x : A, exists ! y : B, R' x y

subrelation (fun (x : A) (y : B) => P x /\ R' x y) R /\ (forall x : A, P x -> exists ! y : B, P x /\ R' x y)
firstorder. Qed.

forall A B : Type, GuardedRelationalChoice_on A B -> RelationalChoice_on A B

forall A B : Type, GuardedRelationalChoice_on A B -> RelationalChoice_on A B
A, B:Type
GAC_rel:GuardedRelationalChoice_on A B
R:A -> B -> Prop
H:forall x : A, exists y : B, R x y

exists R' : A -> B -> Prop, subrelation R' R /\ (forall x : A, exists ! y : B, R' x y)
A, B:Type
GAC_rel:GuardedRelationalChoice_on A B
R:A -> B -> Prop
H:forall x : A, exists y : B, R x y

forall x : A, True -> exists y : B, R x y
A, B:Type
GAC_rel:GuardedRelationalChoice_on A B
R:A -> B -> Prop
H:forall x : A, exists y : B, R x y
R':A -> B -> Prop
HR'R:subrelation R' R
H0:forall x : A, True -> exists ! y : B, R' x y
exists R'0 : A -> B -> Prop, subrelation R'0 R /\ (forall x : A, exists ! y : B, R'0 x y)
A, B:Type
GAC_rel:GuardedRelationalChoice_on A B
R:A -> B -> Prop
H:forall x : A, exists y : B, R x y
R':A -> B -> Prop
HR'R:subrelation R' R
H0:forall x : A, True -> exists ! y : B, R' x y

exists R'0 : A -> B -> Prop, subrelation R'0 R /\ (forall x : A, exists ! y : B, R'0 x y)
exists R'; firstorder. Qed.

ProofIrrelevance -> GuardedRelationalChoice <-> RelationalChoice

ProofIrrelevance -> GuardedRelationalChoice <-> RelationalChoice
intuition auto using guarded_rel_choice_imp_rel_choice, rel_choice_and_proof_irrel_imp_guarded_rel_choice. Qed.
OAC_rel = GAC_rel

GuardedRelationalChoice <-> OmniscientRelationalChoice

GuardedRelationalChoice <-> OmniscientRelationalChoice

GuardedRelationalChoice -> OmniscientRelationalChoice

OmniscientRelationalChoice -> GuardedRelationalChoice
GAC_rel:GuardedRelationalChoice
A, B:Type
R:A -> B -> Prop

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

OmniscientRelationalChoice -> GuardedRelationalChoice

OmniscientRelationalChoice -> GuardedRelationalChoice
OAC_rel:OmniscientRelationalChoice
A, B:Type
P:A -> Prop
R:A -> B -> Prop
H:forall x : A, P x -> exists y : B, R x y

exists R' : A -> B -> Prop, subrelation R' R /\ (forall x : A, P x -> exists ! y : B, R' x y)
destruct (OAC_rel A B R) as (f,Hf); exists f; firstorder. Qed. (**********************************************************************)

AC_fun + IGP = GAC_fun = OAC_fun = AC_fun + Drinker

AC_fun + IGP = GAC_fun

GuardedFunctionalChoice -> IndependenceOfGeneralPremises

GuardedFunctionalChoice -> IndependenceOfGeneralPremises
GAC_fun:GuardedFunctionalChoice
A:Type
P:A -> Prop
Q:Prop
Inh:inhabited A
H:Q -> exists x : A, P x

exists x : A, Q -> P x
GAC_fun:GuardedFunctionalChoice
A:Type
P:A -> Prop
Q:Prop
Inh:inhabited A
H:Q -> exists x : A, P x

unit -> Q -> exists y : A, P y
GAC_fun:GuardedFunctionalChoice
A:Type
P:A -> Prop
Q:Prop
Inh:inhabited A
H:Q -> exists x : A, P x
f:unit -> A
Hf:forall x : unit, Q -> P (f x)
exists x : A, Q -> P x
GAC_fun:GuardedFunctionalChoice
A:Type
P:A -> Prop
Q:Prop
Inh:inhabited A
H:Q -> exists x : A, P x
f:unit -> A
Hf:forall x : unit, Q -> P (f x)

exists x : A, Q -> P x
exists (f tt); auto. Qed.

GuardedFunctionalChoice -> FunctionalChoiceOnInhabitedSet

GuardedFunctionalChoice -> FunctionalChoiceOnInhabitedSet
GAC_fun:GuardedFunctionalChoice
A, B:Type
Inh:inhabited B
R:A -> B -> Prop
H:forall x : A, exists y : B, R x y

exists f : A -> B, forall x : A, R x (f x)
GAC_fun:GuardedFunctionalChoice
A, B:Type
Inh:inhabited B
R:A -> B -> Prop
H:forall x : A, exists y : B, R x y

forall x : A, True -> exists y : B, R x y
GAC_fun:GuardedFunctionalChoice
A, B:Type
Inh:inhabited B
R:A -> B -> Prop
H:forall x : A, exists y : B, R x y
f:A -> B
Hf:forall x : A, True -> R x (f x)
exists f0 : A -> B, forall x : A, R x (f0 x)
GAC_fun:GuardedFunctionalChoice
A, B:Type
Inh:inhabited B
R:A -> B -> Prop
H:forall x : A, exists y : B, R x y
f:A -> B
Hf:forall x : A, True -> R x (f x)

exists f0 : A -> B, forall x : A, R x (f0 x)
exists f; auto. Qed.

FunctionalChoiceOnInhabitedSet -> IndependenceOfGeneralPremises -> GuardedFunctionalChoice

FunctionalChoiceOnInhabitedSet -> IndependenceOfGeneralPremises -> GuardedFunctionalChoice
AC_fun:FunctionalChoiceOnInhabitedSet
IndPrem:IndependenceOfGeneralPremises
A, B:Type
P:A -> Prop
R:A -> B -> Prop
Inh:inhabited B
H:forall x : A, P x -> exists y : B, R x y

exists f : A -> B, forall x : A, P x -> R x (f x)
AC_fun:FunctionalChoiceOnInhabitedSet
IndPrem:IndependenceOfGeneralPremises
A, B:Type
P:A -> Prop
R:A -> B -> Prop
Inh:inhabited B
H:forall x : A, P x -> exists y : B, R x y

forall x : A, exists y : B, P x -> R x y
intro x; apply IndPrem; eauto. Qed.

FunctionalChoiceOnInhabitedSet /\ IndependenceOfGeneralPremises <-> GuardedFunctionalChoice

FunctionalChoiceOnInhabitedSet /\ IndependenceOfGeneralPremises <-> GuardedFunctionalChoice
intuition auto using guarded_fun_choice_imp_indep_of_general_premises, guarded_fun_choice_imp_fun_choice, fun_choice_and_indep_general_prem_imp_guarded_fun_choice. Qed.
AC_fun + Drinker = OAC_fun
This was already observed by Bell [Bell]

OmniscientFunctionalChoice -> SmallDrinker'sParadox

OmniscientFunctionalChoice -> SmallDrinker'sParadox
OAC_fun:OmniscientFunctionalChoice
A:Type
P:A -> Prop
Inh:inhabited A

exists x : A, (exists x0 : A, P x0) -> P x
OAC_fun:OmniscientFunctionalChoice
A:Type
P:A -> Prop
Inh:inhabited A

inhabited A
OAC_fun:OmniscientFunctionalChoice
A:Type
P:A -> Prop
Inh:inhabited A
f:unit -> A
Hf:forall x : unit, (exists y : A, P y) -> P (f x)
exists x : A, (exists x0 : A, P x0) -> P x
OAC_fun:OmniscientFunctionalChoice
A:Type
P:A -> Prop
Inh:inhabited A
f:unit -> A
Hf:forall x : unit, (exists y : A, P y) -> P (f x)

exists x : A, (exists x0 : A, P x0) -> P x
exists (f tt); firstorder. Qed.

OmniscientFunctionalChoice -> FunctionalChoiceOnInhabitedSet

OmniscientFunctionalChoice -> FunctionalChoiceOnInhabitedSet
OAC_fun:OmniscientFunctionalChoice
A, B:Type
Inh:inhabited B
R:A -> B -> Prop
H:forall x : A, exists y : B, R x y

exists f : A -> B, forall x : A, R x (f x)
OAC_fun:OmniscientFunctionalChoice
A, B:Type
Inh:inhabited B
R:A -> B -> Prop
H:forall x : A, exists y : B, R x y
f:A -> B
Hf:forall x : A, (exists y : B, R x y) -> R x (f x)

exists f0 : A -> B, forall x : A, R x (f0 x)
exists f; firstorder. Qed.

FunctionalChoiceOnInhabitedSet -> SmallDrinker'sParadox -> OmniscientFunctionalChoice

FunctionalChoiceOnInhabitedSet -> SmallDrinker'sParadox -> OmniscientFunctionalChoice
AC_fun:FunctionalChoiceOnInhabitedSet
Drinker:SmallDrinker'sParadox
A, B:Type
R:A -> B -> Prop
Inh:inhabited B

exists f : A -> B, forall x : A, (exists y : B, R x y) -> R x (f x)
AC_fun:FunctionalChoiceOnInhabitedSet
Drinker:SmallDrinker'sParadox
A, B:Type
R:A -> B -> Prop
Inh:inhabited B

forall x : A, exists y : B, (exists y0 : B, R x y0) -> R x y
AC_fun:FunctionalChoiceOnInhabitedSet
Drinker:SmallDrinker'sParadox
A, B:Type
R:A -> B -> Prop
Inh:inhabited B
f:A -> B
Hf:forall x : A, (exists y : B, R x y) -> R x (f x)
exists f0 : A -> B, forall x : A, (exists y : B, R x y) -> R x (f0 x)
AC_fun:FunctionalChoiceOnInhabitedSet
Drinker:SmallDrinker'sParadox
A, B:Type
R:A -> B -> Prop
Inh:inhabited B
f:A -> B
Hf:forall x : A, (exists y : B, R x y) -> R x (f x)

exists f0 : A -> B, forall x : A, (exists y : B, R x y) -> R x (f0 x)
exists f; assumption. Qed.

FunctionalChoiceOnInhabitedSet /\ SmallDrinker'sParadox <-> OmniscientFunctionalChoice

FunctionalChoiceOnInhabitedSet /\ SmallDrinker'sParadox <-> OmniscientFunctionalChoice
intuition auto using omniscient_fun_choice_imp_small_drinker, omniscient_fun_choice_imp_fun_choice, fun_choice_and_small_drinker_imp_omniscient_fun_choice. Qed.
OAC_fun = GAC_fun
This is derivable from the intuitionistic equivalence between IGP and Drinker but we give a direct proof

GuardedFunctionalChoice <-> OmniscientFunctionalChoice

GuardedFunctionalChoice <-> OmniscientFunctionalChoice

GuardedFunctionalChoice -> OmniscientFunctionalChoice

OmniscientFunctionalChoice -> GuardedFunctionalChoice
GAC_fun:GuardedFunctionalChoice
A, B:Type
R:A -> B -> Prop
Inh:inhabited B

exists f : A -> B, forall x : A, (exists y : B, R x y) -> R x (f x)

OmniscientFunctionalChoice -> GuardedFunctionalChoice

OmniscientFunctionalChoice -> GuardedFunctionalChoice
OAC_fun:OmniscientFunctionalChoice
A, B:Type
P:A -> Prop
R:A -> B -> Prop
Inh:inhabited B
H: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_fun:OmniscientFunctionalChoice
A, B:Type
P:A -> Prop
R:A -> B -> Prop
Inh:inhabited B
H:forall x : A, P x -> exists y : B, R x y
f:A -> B
Hf:forall x : A, (exists y : B, R x y) -> R x (f x)

exists f0 : A -> B, forall x : A, P x -> R x (f0 x)
exists f; firstorder. Qed. (**********************************************************************)

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

D_iota -> ID_iota

IotaStatement -> ConstructiveDefiniteDescription

IotaStatement -> ConstructiveDefiniteDescription
D_iota:IotaStatement
A:Type
P:A -> Prop
H:exists ! x : A, P x

{x : A | P x}
D_iota:IotaStatement
A:Type
P:A -> Prop
H:exists ! x : A, P x

inhabited A
D_iota:IotaStatement
A:Type
P:A -> Prop
H:exists ! x0 : A, P x0
x:A
H1:(exists ! x0 : A, P x0) -> P x
{x0 : A | P x0}
D_iota:IotaStatement
A:Type
P:A -> Prop
H:exists ! x0 : A, P x0
x:A
H1:(exists ! x0 : A, P x0) -> P x

{x0 : A | P x0}
exists x; apply H1; assumption. Qed.
ID_epsilon + Drinker <-> D_epsilon

EpsilonStatement -> ConstructiveIndefiniteDescription

EpsilonStatement -> ConstructiveIndefiniteDescription
D_epsilon:EpsilonStatement
A:Type
P:A -> Prop
H:exists x : A, P x

{x : A | P x}
D_epsilon:EpsilonStatement
A:Type
P:A -> Prop
H:exists x : A, P x

inhabited A
D_epsilon:EpsilonStatement
A:Type
P:A -> Prop
H:exists x0 : A, P x0
x:A
H1:(exists x0 : A, P x0) -> P x
{x0 : A | P x0}
D_epsilon:EpsilonStatement
A:Type
P:A -> Prop
H:exists x0 : A, P x0
x:A
H1:(exists x0 : A, P x0) -> P x

{x0 : A | P x0}
exists x; apply H1; assumption. Qed.

SmallDrinker'sParadox -> ConstructiveIndefiniteDescription -> EpsilonStatement

SmallDrinker'sParadox -> ConstructiveIndefiniteDescription -> EpsilonStatement
intros Drinkers D_epsilon A P Inh; apply D_epsilon; apply Drinkers; assumption. Qed.

EpsilonStatement -> SmallDrinker'sParadox

EpsilonStatement -> SmallDrinker'sParadox
intros D_epsilon A P Inh; edestruct D_epsilon; eauto. Qed.

((SmallDrinker'sParadox * ConstructiveIndefiniteDescription -> EpsilonStatement) * (EpsilonStatement -> SmallDrinker'sParadox * ConstructiveIndefiniteDescription))%type

((SmallDrinker'sParadox * ConstructiveIndefiniteDescription -> EpsilonStatement) * (EpsilonStatement -> SmallDrinker'sParadox * ConstructiveIndefiniteDescription))%type
intuition auto using epsilon_imp_constructive_indefinite_description, constructive_indefinite_description_and_small_drinker_imp_epsilon, epsilon_imp_small_drinker. Qed. (**********************************************************************)

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 relation with nat as codomain
Require Import Wf_nat.
Require Import Decidable.


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

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

forall R : A -> nat -> Prop, (forall (x : A) (y : nat), decidable (R x y)) -> FunctionalChoice_on_rel R
A:Type
Descr:FunctionalRelReification_on A nat
R:A -> nat -> Prop
Rdec:forall (x : A) (y : nat), decidable (R x y)
H:forall x : A, exists y : nat, R x y

exists f : A -> nat, forall x : A, R x (f x)
A:Type
Descr:FunctionalRelReification_on A nat
R:A -> nat -> Prop
Rdec:forall (x : A) (y : nat), decidable (R x y)
H:forall x : A, exists y : nat, R x y
R':=fun (x : A) (y : nat) => R x y /\ (forall y' : nat, R x y' -> y <= y'):A -> nat -> Prop

exists f : A -> nat, forall x : A, R x (f x)
A:Type
Descr:FunctionalRelReification_on A nat
R:A -> nat -> Prop
Rdec:forall (x : A) (y : nat), decidable (R x y)
H:forall x : A, exists y : nat, R x y
R':=fun (x : A) (y : nat) => R x y /\ (forall y' : nat, R x y' -> y <= y'):A -> nat -> Prop

forall x : A, exists ! y : nat, R' x y
A:Type
Descr:FunctionalRelReification_on A nat
R:A -> nat -> Prop
Rdec:forall (x : A) (y : nat), decidable (R x y)
H:forall x : A, exists y : nat, R x y
R':=fun (x : A) (y : nat) => R x y /\ (forall y' : nat, R x y' -> y <= y'):A -> nat -> Prop
f:A -> nat
Hf:forall x : A, R' x (f x)
exists f0 : A -> nat, forall x : A, R x (f0 x)
A:Type
Descr:FunctionalRelReification_on A nat
R:A -> nat -> Prop
Rdec:forall (x0 : A) (y : nat), decidable (R x0 y)
H:forall x0 : A, exists y : nat, R x0 y
R':=fun (x0 : A) (y : nat) => R x0 y /\ (forall y' : nat, R x0 y' -> y <= y'):A -> nat -> Prop
x:A

exists ! y : nat, R' x y
A:Type
Descr:FunctionalRelReification_on A nat
R:A -> nat -> Prop
Rdec:forall (x : A) (y : nat), decidable (R x y)
H:forall x : A, exists y : nat, R x y
R':=fun (x : A) (y : nat) => R x y /\ (forall y' : nat, R x y' -> y <= y'):A -> nat -> Prop
f:A -> nat
Hf:forall x : A, R' x (f x)
exists f0 : A -> nat, forall x : A, R x (f0 x)
A:Type
Descr:FunctionalRelReification_on A nat
R:A -> nat -> Prop
Rdec:forall (x0 : A) (y : nat), decidable (R x0 y)
H:forall x0 : A, exists y : nat, R x0 y
R':=fun (x0 : A) (y : nat) => R x0 y /\ (forall y' : nat, R x0 y' -> y <= y'):A -> nat -> Prop
x:A

forall n : nat, R x n \/ ~ R x n
A:Type
Descr:FunctionalRelReification_on A nat
R:A -> nat -> Prop
Rdec:forall (x0 : A) (y : nat), decidable (R x0 y)
H:forall x0 : A, exists y : nat, R x0 y
R':=fun (x0 : A) (y : nat) => R x0 y /\ (forall y' : nat, R x0 y' -> y <= y'):A -> nat -> Prop
x:A
exists n : nat, R x n
A:Type
Descr:FunctionalRelReification_on A nat
R:A -> nat -> Prop
Rdec:forall (x : A) (y : nat), decidable (R x y)
H:forall x : A, exists y : nat, R x y
R':=fun (x : A) (y : nat) => R x y /\ (forall y' : nat, R x y' -> y <= y'):A -> nat -> Prop
f:A -> nat
Hf:forall x : A, R' x (f x)
exists f0 : A -> nat, forall x : A, R x (f0 x)
A:Type
Descr:FunctionalRelReification_on A nat
R:A -> nat -> Prop
Rdec:forall (x0 : A) (y : nat), decidable (R x0 y)
H:forall x0 : A, exists y : nat, R x0 y
R':=fun (x0 : A) (y : nat) => R x0 y /\ (forall y' : nat, R x0 y' -> y <= y'):A -> nat -> Prop
x:A

exists n : nat, R x n
A:Type
Descr:FunctionalRelReification_on A nat
R:A -> nat -> Prop
Rdec:forall (x : A) (y : nat), decidable (R x y)
H:forall x : A, exists y : nat, R x y
R':=fun (x : A) (y : nat) => R x y /\ (forall y' : nat, R x y' -> y <= y'):A -> nat -> Prop
f:A -> nat
Hf:forall x : A, R' x (f x)
exists f0 : A -> nat, forall x : A, R x (f0 x)
A:Type
Descr:FunctionalRelReification_on A nat
R:A -> nat -> Prop
Rdec:forall (x : A) (y : nat), decidable (R x y)
H:forall x : A, exists y : nat, R x y
R':=fun (x : A) (y : nat) => R x y /\ (forall y' : nat, R x y' -> y <= y'):A -> nat -> Prop
f:A -> nat
Hf:forall x : A, R' x (f x)

exists f0 : A -> nat, forall x : A, R x (f0 x)
A:Type
Descr:FunctionalRelReification_on A nat
R:A -> nat -> Prop
Rdec:forall (x : A) (y : nat), decidable (R x y)
H:forall x : A, exists y : nat, R x y
R':=fun (x : A) (y : nat) => R x y /\ (forall y' : nat, R x y' -> y <= y'):A -> nat -> Prop
f:A -> nat
Hf:forall x : A, R' x (f x)

forall x : A, R x (f x)
A:Type
Descr:FunctionalRelReification_on A nat
R:A -> nat -> Prop
Rdec:forall (x0 : A) (y : nat), decidable (R x0 y)
H:forall x0 : A, exists y : nat, R x0 y
R':=fun (x0 : A) (y : nat) => R x0 y /\ (forall y' : nat, R x0 y' -> y <= y'):A -> nat -> Prop
f:A -> nat
Hf:forall x0 : A, R' x0 (f x0)
x:A

R x (f x)
A:Type
Descr:FunctionalRelReification_on A nat
R:A -> nat -> Prop
Rdec:forall (x0 : A) (y : nat), decidable (R x0 y)
H:forall x0 : A, exists y : nat, R x0 y
R':=fun (x0 : A) (y : nat) => R x0 y /\ (forall y' : nat, R x0 y' -> y <= y'):A -> nat -> Prop
f:A -> nat
Hf:forall x0 : A, R' x0 (f x0)
x:A
Hfx:R x (f x)

R x (f x)
assumption. Qed. (**********************************************************************)

AC_fun = AC_fun_dep = AC_trunc

Choice on dependent and non dependent function types are equivalent

The easy part

DependentFunctionalChoice -> FunctionalChoice

DependentFunctionalChoice -> FunctionalChoice
AC_depfun:DependentFunctionalChoice
A, B:Type
R:A -> B -> Prop
H:forall x : A, exists y : B, R x y

exists f : A -> B, forall x : A, R x (f x)
AC_depfun:DependentFunctionalChoice
A, B:Type
R:A -> B -> Prop
H:forall x : A, exists y : B, R x y
f:A -> B
Hf:forall x : A, R x (f x)

exists f0 : A -> B, forall x : A, R x (f0 x)
exists f; trivial. Qed.
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.


FunctionalChoice -> DependentFunctionalChoice

FunctionalChoice -> DependentFunctionalChoice
AC_fun:FunctionalChoice
A:Type
B:A -> Type
R:forall x : A, B x -> Prop
H: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:FunctionalChoice
A:Type
B:A -> Type
R:forall x : A, B x -> Prop
H:forall x : A, exists y : B x, R x y
B':={x : A & B x}:Type

exists f : forall x : A, B x, forall x : A, R x (f x)
AC_fun:FunctionalChoice
A:Type
B:A -> Type
R:forall x : A, B x -> Prop
H:forall x : A, exists y : B x, R x y
B':={x : A & B x}:Type
R':=fun (x : A) (y : B') => projT1 y = x /\ R (projT1 y) (projT2 y):A -> B' -> Prop

exists f : forall x : A, B x, forall x : A, R x (f x)
AC_fun:FunctionalChoice
A:Type
B:A -> Type
R:forall x : A, B x -> Prop
H:forall x : A, exists y : B x, R x y
B':={x : A & B x}:Type
R':=fun (x : A) (y : B') => projT1 y = x /\ R (projT1 y) (projT2 y):A -> B' -> Prop

forall x : A, exists y : B', R' x y
AC_fun:FunctionalChoice
A:Type
B:A -> Type
R:forall x : A, B x -> Prop
H:forall x : A, exists y : B x, R x y
B':={x : A & B x}:Type
R':=fun (x : A) (y : B') => projT1 y = x /\ R (projT1 y) (projT2 y):A -> B' -> Prop
f:A -> B'
Hf:forall x : A, R' x (f x)
exists f0 : forall x : A, B x, forall x : A, R x (f0 x)
AC_fun:FunctionalChoice
A:Type
B:A -> Type
R:forall x0 : A, B x0 -> Prop
H:forall x0 : A, exists y : B x0, R x0 y
B':={x0 : A & B x0}:Type
R':=fun (x0 : A) (y : B') => projT1 y = x0 /\ R (projT1 y) (projT2 y):A -> B' -> Prop
x:A

exists y : B', R' x y
AC_fun:FunctionalChoice
A:Type
B:A -> Type
R:forall x : A, B x -> Prop
H:forall x : A, exists y : B x, R x y
B':={x : A & B x}:Type
R':=fun (x : A) (y : B') => projT1 y = x /\ R (projT1 y) (projT2 y):A -> B' -> Prop
f:A -> B'
Hf:forall x : A, R' x (f x)
exists f0 : forall x : A, B x, forall x : A, R x (f0 x)
AC_fun:FunctionalChoice
A:Type
B:A -> Type
R:forall x0 : A, B x0 -> Prop
H:forall x0 : A, exists y0 : B x0, R x0 y0
B':={x0 : A & B x0}:Type
R':=fun (x0 : A) (y0 : B') => projT1 y0 = x0 /\ R (projT1 y0) (projT2 y0):A -> B' -> Prop
x:A
y:B x
Hy:R x y

exists y0 : B', R' x y0
AC_fun:FunctionalChoice
A:Type
B:A -> Type
R:forall x : A, B x -> Prop
H:forall x : A, exists y : B x, R x y
B':={x : A & B x}:Type
R':=fun (x : A) (y : B') => projT1 y = x /\ R (projT1 y) (projT2 y):A -> B' -> Prop
f:A -> B'
Hf:forall x : A, R' x (f x)
exists f0 : forall x : A, B x, forall x : A, R x (f0 x)
AC_fun:FunctionalChoice
A:Type
B:A -> Type
R:forall x0 : A, B x0 -> Prop
H:forall x0 : A, exists y0 : B x0, R x0 y0
B':={x0 : A & B x0}:Type
R':=fun (x0 : A) (y0 : B') => projT1 y0 = x0 /\ R (projT1 y0) (projT2 y0):A -> B' -> Prop
x:A
y:B x
Hy:R x y

R' x (existT (fun x0 : A => B x0) x y)
AC_fun:FunctionalChoice
A:Type
B:A -> Type
R:forall x : A, B x -> Prop
H:forall x : A, exists y : B x, R x y
B':={x : A & B x}:Type
R':=fun (x : A) (y : B') => projT1 y = x /\ R (projT1 y) (projT2 y):A -> B' -> Prop
f:A -> B'
Hf:forall x : A, R' x (f x)
exists f0 : forall x : A, B x, forall x : A, R x (f0 x)
AC_fun:FunctionalChoice
A:Type
B:A -> Type
R:forall x : A, B x -> Prop
H:forall x : A, exists y : B x, R x y
B':={x : A & B x}:Type
R':=fun (x : A) (y : B') => projT1 y = x /\ R (projT1 y) (projT2 y):A -> B' -> Prop
f:A -> B'
Hf:forall x : A, R' x (f x)

exists f0 : forall x : A, B x, forall x : A, R x (f0 x)
AC_fun:FunctionalChoice
A:Type
B:A -> Type
R:forall x : A, B x -> Prop
H:forall x : A, exists y : B x, R x y
B':={x : A & B x}:Type
R':=fun (x : A) (y : B') => projT1 y = x /\ R (projT1 y) (projT2 y):A -> B' -> Prop
f:A -> B'
Hf:forall x : A, R' x (f x)

forall x : A, R x (eq_rect (projT1 (f x)) (fun x0 : A => B x0) (projT2 (f x)) x (proj1_inf (Hf x)))
AC_fun:FunctionalChoice
A:Type
B:A -> Type
R:forall x0 : A, B x0 -> Prop
H:forall x0 : A, exists y : B x0, R x0 y
B':={x0 : A & B x0}:Type
R':=fun (x0 : A) (y : B') => projT1 y = x0 /\ R (projT1 y) (projT2 y):A -> B' -> Prop
f:A -> B'
Hf:forall x0 : A, R' x0 (f x0)
x:A
Heq:projT1 (f x) = x
HR:R (projT1 (f x)) (projT2 (f x))

R x (eq_rect (projT1 (f x)) (fun x0 : A => B x0) (projT2 (f x)) x (proj1_inf (conj Heq HR)))
AC_fun:FunctionalChoice
A:Type
B:A -> Type
R:forall x1 : A, B x1 -> Prop
H:forall x1 : A, exists y : B x1, R x1 y
B':={x1 : A & B x1}:Type
R':=fun (x1 : A) (y : B') => projT1 y = x1 /\ R (projT1 y) (projT2 y):A -> B' -> Prop
f:A -> B'
Hf:forall x1 : A, R' x1 (f x1)
x, x0:A
b:B x0
Heq:x0 = x
HR:R x0 b

R x (eq_rect x0 (fun x1 : A => B x1) b x Heq)
destruct Heq using eq_indd; trivial. Qed.

Functional choice and truncation choice are equivalent


FunctionalChoice -> InhabitedForallCommute

FunctionalChoice -> InhabitedForallCommute
choose0:FunctionalChoice
A:Type
B:A -> Type
Hinhab:forall x : A, inhabited (B x)

inhabited (forall x : A, B x)
A:Type
B:A -> Type
Hinhab:forall x : A, inhabited (B x)
choose:DependentFunctionalChoice

inhabited (forall x : A, B x)
A:Type
B:A -> Type
Hinhab:forall x : A, inhabited (B x)
choose:DependentFunctionalChoice

forall x : A, exists _ : B x, True
A:Type
B:A -> Type
Hinhab:forall x : A, inhabited (B x)
choose:DependentFunctionalChoice
Hexists:forall x : A, exists _ : B x, True
inhabited (forall x : A, B x)
A:Type
B:A -> Type
Hinhab:forall x : A, inhabited (B x)
choose:DependentFunctionalChoice

forall x : A, exists _ : B x, True
A:Type
B:A -> Type
Hinhab:forall x0 : A, inhabited (B x0)
choose:DependentFunctionalChoice
x:A

inhabited {_ : B x | True}
A:Type
B:A -> Type
Hinhab:forall x0 : A, inhabited (B x0)
choose:DependentFunctionalChoice
x:A

B x -> {_ : B x | True}
intros y;exists y;exact I.
A:Type
B:A -> Type
Hinhab:forall x : A, inhabited (B x)
choose:DependentFunctionalChoice
Hexists:forall x : A, exists _ : B x, True

inhabited (forall x : A, B x)
A:Type
B:A -> Type
Hinhab:forall x : A, inhabited (B x)
choose:DependentFunctionalChoice
Hexists:exists _ : forall x : A, B x, A -> True

inhabited (forall x : A, B x)
A:Type
B:A -> Type
Hinhab:forall x : A, inhabited (B x)
choose:DependentFunctionalChoice
f:forall x : A, B x

inhabited (forall x : A, B x)
exact (inhabits f). Qed.

InhabitedForallCommute -> FunctionalChoice

InhabitedForallCommute -> FunctionalChoice
choose:InhabitedForallCommute
A, B:Type
R:A -> B -> Prop
Hexists:forall x : A, exists y : B, R x y

exists f : A -> B, forall x : A, R x (f x)
choose:InhabitedForallCommute
A, B:Type
R:A -> B -> Prop
Hexists:forall x : A, exists y : B, R x y

forall x : A, inhabited {y : B | R x y}
choose:InhabitedForallCommute
A, B:Type
R:A -> B -> Prop
Hexists:forall x : A, exists y : B, R x y
Hinhab:forall x : A, inhabited {y : B | R x y}
exists f : A -> B, forall x : A, R x (f x)
choose:InhabitedForallCommute
A, B:Type
R:A -> B -> Prop
Hexists:forall x : A, exists y : B, R x y

forall x : A, inhabited {y : B | R x y}
intros x;apply exists_to_inhabited_sig;trivial.
choose:InhabitedForallCommute
A, B:Type
R:A -> B -> Prop
Hexists:forall x : A, exists y : B, R x y
Hinhab:forall x : A, inhabited {y : B | R x y}

exists f : A -> B, forall x : A, R x (f x)
choose:InhabitedForallCommute
A, B:Type
R:A -> B -> Prop
Hexists:forall x : A, exists y : B, R x y
Hinhab:inhabited (forall x : A, {y : B | R x y})

exists f : A -> B, forall x : A, R x (f x)
choose:InhabitedForallCommute
A, B:Type
R:A -> B -> Prop
Hexists:forall x : A, exists y : B, R x y
f:forall x : A, {y : B | R x y}

exists f0 : A -> B, forall x : A, R x (f0 x)
choose:InhabitedForallCommute
A, B:Type
R:A -> B -> Prop
Hexists:forall x : A, exists y : B, R x y
f:forall x : A, {y : B | R x y}

forall x : A, R x (proj1_sig (f x))
exact (fun x => proj2_sig (f x)). Qed.

Reification of dependent and non dependent functional relation are equivalent

The easy part

DependentFunctionalRelReification -> FunctionalRelReification

DependentFunctionalRelReification -> FunctionalRelReification
DepFunReify:DependentFunctionalRelReification
A, B:Type
R:A -> B -> Prop
H:forall x : A, exists ! y : B, R x y

exists f : A -> B, forall x : A, R x (f x)
DepFunReify:DependentFunctionalRelReification
A, B:Type
R:A -> B -> Prop
H:forall x : A, exists ! y : B, R x y
f:A -> B
Hf:forall x : A, R x (f x)

exists f0 : A -> B, forall x : A, R x (f0 x)
exists f; trivial. Qed.
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

FunctionalRelReification -> DependentFunctionalRelReification

FunctionalRelReification -> DependentFunctionalRelReification
AC_fun:FunctionalRelReification
A:Type
B:A -> Type
R:forall x : A, B x -> Prop
H: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:FunctionalRelReification
A:Type
B:A -> Type
R:forall x : A, B x -> Prop
H:forall x : A, exists ! y : B x, R x y
B':={x : A & B x}:Type

exists f : forall x : A, B x, forall x : A, R x (f x)
AC_fun:FunctionalRelReification
A:Type
B:A -> Type
R:forall x : A, B x -> Prop
H:forall x : A, exists ! y : B x, R x y
B':={x : A & B x}:Type
R':=fun (x : A) (y : B') => projT1 y = x /\ R (projT1 y) (projT2 y):A -> B' -> Prop

exists f : forall x : A, B x, forall x : A, R x (f x)
AC_fun:FunctionalRelReification
A:Type
B:A -> Type
R:forall x : A, B x -> Prop
H:forall x : A, exists ! y : B x, R x y
B':={x : A & B x}:Type
R':=fun (x : A) (y : B') => projT1 y = x /\ R (projT1 y) (projT2 y):A -> B' -> Prop

forall x : A, exists ! y : B', R' x y
AC_fun:FunctionalRelReification
A:Type
B:A -> Type
R:forall x : A, B x -> Prop
H:forall x : A, exists ! y : B x, R x y
B':={x : A & B x}:Type
R':=fun (x : A) (y : B') => projT1 y = x /\ R (projT1 y) (projT2 y):A -> B' -> Prop
f:A -> B'
Hf:forall x : A, R' x (f x)
exists f0 : forall x : A, B x, forall x : A, R x (f0 x)
AC_fun:FunctionalRelReification
A:Type
B:A -> Type
R:forall x0 : A, B x0 -> Prop
H:forall x0 : A, exists ! y : B x0, R x0 y
B':={x0 : A & B x0}:Type
R':=fun (x0 : A) (y : B') => projT1 y = x0 /\ R (projT1 y) (projT2 y):A -> B' -> Prop
x:A

exists ! y : B', R' x y
AC_fun:FunctionalRelReification
A:Type
B:A -> Type
R:forall x : A, B x -> Prop
H:forall x : A, exists ! y : B x, R x y
B':={x : A & B x}:Type
R':=fun (x : A) (y : B') => projT1 y = x /\ R (projT1 y) (projT2 y):A -> B' -> Prop
f:A -> B'
Hf:forall x : A, R' x (f x)
exists f0 : forall x : A, B x, forall x : A, R x (f0 x)
AC_fun:FunctionalRelReification
A:Type
B:A -> Type
R:forall x0 : A, B x0 -> Prop
H:forall x0 : A, exists ! y0 : B x0, R x0 y0
B':={x0 : A & B x0}:Type
R':=fun (x0 : A) (y0 : B') => projT1 y0 = x0 /\ R (projT1 y0) (projT2 y0):A -> B' -> Prop
x:A
y:B x
Hy:R x y
Huni:forall x' : B x, R x x' -> y = x'

exists ! y0 : B', R' x y0
AC_fun:FunctionalRelReification
A:Type
B:A -> Type
R:forall x : A, B x -> Prop
H:forall x : A, exists ! y : B x, R x y
B':={x : A & B x}:Type
R':=fun (x : A) (y : B') => projT1 y = x /\ R (projT1 y) (projT2 y):A -> B' -> Prop
f:A -> B'
Hf:forall x : A, R' x (f x)
exists f0 : forall x : A, B x, forall x : A, R x (f0 x)
AC_fun:FunctionalRelReification
A:Type
B:A -> Type
R:forall x0 : A, B x0 -> Prop
H:forall x0 : A, exists ! y0 : B x0, R x0 y0
B':={x0 : A & B x0}:Type
R':=fun (x0 : A) (y0 : B') => projT1 y0 = x0 /\ R (projT1 y0) (projT2 y0):A -> B' -> Prop
x:A
y:B x
Hy:R x y
Huni:forall x' : B x, R x x' -> y = x'

unique (fun y0 : B' => R' x y0) (existT (fun x0 : A => B x0) x y)
AC_fun:FunctionalRelReification
A:Type
B:A -> Type
R:forall x : A, B x -> Prop
H:forall x : A, exists ! y : B x, R x y
B':={x : A & B x}:Type
R':=fun (x : A) (y : B') => projT1 y = x /\ R (projT1 y) (projT2 y):A -> B' -> Prop
f:A -> B'
Hf:forall x : A, R' x (f x)
exists f0 : forall x : A, B x, forall x : A, R x (f0 x)
AC_fun:FunctionalRelReification
A:Type
B:A -> Type
R:forall x0 : A, B x0 -> Prop
H:forall x0 : A, exists ! y0 : B x0, R x0 y0
B':={x0 : A & B x0}:Type
R':=fun (x0 : A) (y0 : B') => projT1 y0 = x0 /\ R (projT1 y0) (projT2 y0):A -> B' -> Prop
x:A
y:B x
Hy:R x y
Huni:forall x' : B x, R x x' -> y = x'

forall x' : B', R' x x' -> existT (fun x0 : A => B x0) x y = x'
AC_fun:FunctionalRelReification
A:Type
B:A -> Type
R:forall x : A, B x -> Prop
H:forall x : A, exists ! y : B x, R x y
B':={x : A & B x}:Type
R':=fun (x : A) (y : B') => projT1 y = x /\ R (projT1 y) (projT2 y):A -> B' -> Prop
f:A -> B'
Hf:forall x : A, R' x (f x)
exists f0 : forall x : A, B x, forall x : A, R x (f0 x)
AC_fun:FunctionalRelReification
A:Type
B:A -> Type
R:forall x0 : A, B x0 -> Prop
H:forall x0 : A, exists ! y0 : B x0, R x0 y0
B':={x0 : A & B x0}:Type
R':=fun (x0 : A) (y0 : B') => projT1 y0 = x0 /\ R (projT1 y0) (projT2 y0):A -> B' -> Prop
x:A
y:B x
Hy:R x y
Huni:forall x'0 : B x, R x x'0 -> y = x'0
x':A
y':B x'
Heqx':projT1 (existT (fun x0 : A => B x0) x' y') = x
Hy':R (projT1 (existT (fun x0 : A => B x0) x' y')) (projT2 (existT (fun x0 : A => B x0) x' y'))

existT (fun x0 : A => B x0) x y = existT (fun x0 : A => B x0) x' y'
AC_fun:FunctionalRelReification
A:Type
B:A -> Type
R:forall x : A, B x -> Prop
H:forall x : A, exists ! y : B x, R x y
B':={x : A & B x}:Type
R':=fun (x : A) (y : B') => projT1 y = x /\ R (projT1 y) (projT2 y):A -> B' -> Prop
f:A -> B'
Hf:forall x : A, R' x (f x)
exists f0 : forall x : A, B x, forall x : A, R x (f0 x)
AC_fun:FunctionalRelReification
A:Type
B:A -> Type
R:forall x0 : A, B x0 -> Prop
H:forall x0 : A, exists ! y0 : B x0, R x0 y0
B':={x0 : A & B x0}:Type
R':=fun (x0 : A) (y0 : B') => projT1 y0 = x0 /\ R (projT1 y0) (projT2 y0):A -> B' -> Prop
x:A
y:B x
Hy:R x y
Huni:forall x'0 : B x, R x x'0 -> y = x'0
x':A
y':B x'
Heqx':x' = x
Hy':R x' y'

existT (fun x0 : A => B x0) x y = existT (fun x0 : A => B x0) x' y'
AC_fun:FunctionalRelReification
A:Type
B:A -> Type
R:forall x : A, B x -> Prop
H:forall x : A, exists ! y : B x, R x y
B':={x : A & B x}:Type
R':=fun (x : A) (y : B') => projT1 y = x /\ R (projT1 y) (projT2 y):A -> B' -> Prop
f:A -> B'
Hf:forall x : A, R' x (f x)
exists f0 : forall x : A, B x, forall x : A, R x (f0 x)
AC_fun:FunctionalRelReification
A:Type
B:A -> Type
R:forall x : A, B x -> Prop
H:forall x : A, exists ! y0 : B x, R x y0
B':={x : A & B x}:Type
R':=fun (x : A) (y0 : B') => projT1 y0 = x /\ R (projT1 y0) (projT2 y0):A -> B' -> Prop
x':A
y:B x'
Hy:R x' y
Huni:forall x'0 : B x', R x' x'0 -> y = x'0
y':B x'
Hy':R x' y'

existT (fun x : A => B x) x' y = existT (fun x : A => B x) x' y'
AC_fun:FunctionalRelReification
A:Type
B:A -> Type
R:forall x : A, B x -> Prop
H:forall x : A, exists ! y : B x, R x y
B':={x : A & B x}:Type
R':=fun (x : A) (y : B') => projT1 y = x /\ R (projT1 y) (projT2 y):A -> B' -> Prop
f:A -> B'
Hf:forall x : A, R' x (f x)
exists f0 : forall x : A, B x, forall x : A, R x (f0 x)
AC_fun:FunctionalRelReification
A:Type
B:A -> Type
R:forall x : A, B x -> Prop
H:forall x : A, exists ! y : B x, R x y
B':={x : A & B x}:Type
R':=fun (x : A) (y : B') => projT1 y = x /\ R (projT1 y) (projT2 y):A -> B' -> Prop
f:A -> B'
Hf:forall x : A, R' x (f x)

exists f0 : forall x : A, B x, forall x : A, R x (f0 x)
AC_fun:FunctionalRelReification
A:Type
B:A -> Type
R:forall x : A, B x -> Prop
H:forall x : A, exists ! y : B x, R x y
B':={x : A & B x}:Type
R':=fun (x : A) (y : B') => projT1 y = x /\ R (projT1 y) (projT2 y):A -> B' -> Prop
f:A -> B'
Hf:forall x : A, R' x (f x)

forall x : A, R x (eq_rect (projT1 (f x)) (fun x0 : A => B x0) (projT2 (f x)) x (proj1_inf (Hf x)))
AC_fun:FunctionalRelReification
A:Type
B:A -> Type
R:forall x0 : A, B x0 -> Prop
H:forall x0 : A, exists ! y : B x0, R x0 y
B':={x0 : A & B x0}:Type
R':=fun (x0 : A) (y : B') => projT1 y = x0 /\ R (projT1 y) (projT2 y):A -> B' -> Prop
f:A -> B'
Hf:forall x0 : A, R' x0 (f x0)
x:A
Heq:projT1 (f x) = x
HR:R (projT1 (f x)) (projT2 (f x))

R x (eq_rect (projT1 (f x)) (fun x0 : A => B x0) (projT2 (f x)) x (proj1_inf (conj Heq HR)))
AC_fun:FunctionalRelReification
A:Type
B:A -> Type
R:forall x1 : A, B x1 -> Prop
H:forall x1 : A, exists ! y : B x1, R x1 y
B':={x1 : A & B x1}:Type
R':=fun (x1 : A) (y : B') => projT1 y = x1 /\ R (projT1 y) (projT2 y):A -> B' -> Prop
f:A -> B'
Hf:forall x1 : A, R' x1 (f x1)
x, x0:A
b:B x0
Heq:x0 = x
HR:R x0 b

R x (eq_rect x0 (fun x1 : A => B x1) b x Heq)
destruct Heq using eq_indd; trivial. Qed.

FunctionalRelReification <-> DependentFunctionalRelReification

FunctionalRelReification <-> DependentFunctionalRelReification
intuition auto using non_dep_dep_functional_rel_reification, dep_non_dep_functional_rel_reification. Qed. (**********************************************************************)

Non contradiction of constructive descriptions wrt functional axioms of choice

Non contradiction of indefinite description


forall C : Prop, (ConstructiveIndefiniteDescription -> C) -> FunctionalChoice -> C

forall C : Prop, (ConstructiveIndefiniteDescription -> C) -> FunctionalChoice -> C
C:Prop
H:ConstructiveIndefiniteDescription -> C
AC_fun:FunctionalChoice

C
C:Prop
H:ConstructiveIndefiniteDescription -> C
AC_fun:FunctionalChoice
AC_depfun:DependentFunctionalChoice

C
C:Prop
H:ConstructiveIndefiniteDescription -> C
AC_fun:FunctionalChoice
AC_depfun:DependentFunctionalChoice
A0:={A : Type & {P : A -> Prop & exists x : A, P x}}:Type

C
C:Prop
H:ConstructiveIndefiniteDescription -> C
AC_fun:FunctionalChoice
AC_depfun:DependentFunctionalChoice
A0:={A : Type & {P : A -> Prop & exists x : A, P x}}:Type
B0:=fun x : A0 => projT1 x:A0 -> Type

C
C:Prop
H:ConstructiveIndefiniteDescription -> C
AC_fun:FunctionalChoice
AC_depfun:DependentFunctionalChoice
A0:={A : Type & {P : A -> Prop & exists x : A, P x}}:Type
B0:=fun x : A0 => projT1 x:A0 -> Type
R0:=fun (x : A0) (y : B0 x) => projT1 (projT2 x) y:forall x : A0, B0 x -> Prop

C
C:Prop
H:ConstructiveIndefiniteDescription -> C
AC_fun:FunctionalChoice
AC_depfun:DependentFunctionalChoice
A0:={A : Type & {P : A -> Prop & exists x : A, P x}}:Type
B0:=fun x : A0 => projT1 x:A0 -> Type
R0:=fun (x : A0) (y : B0 x) => projT1 (projT2 x) y:forall x : A0, B0 x -> Prop
H0:=fun x : A0 => projT2 (projT2 x):forall x : A0, (fun P : projT1 x -> Prop => exists x0 : projT1 x, P x0) (projT1 (projT2 x))

C
C:Prop
H:ConstructiveIndefiniteDescription -> C
AC_fun:FunctionalChoice
AC_depfun:DependentFunctionalChoice
A0:={A : Type & {P : A -> Prop & exists x : A, P x}}:Type
B0:=fun x : A0 => projT1 x:A0 -> Type
R0:=fun (x : A0) (y : B0 x) => projT1 (projT2 x) y:forall x : A0, B0 x -> Prop
H0:=fun x : A0 => projT2 (projT2 x):forall x : A0, (fun P : projT1 x -> Prop => exists x0 : projT1 x, P x0) (projT1 (projT2 x))
f:forall x : A0, B0 x
Hf:forall x : A0, R0 x (f x)

C
C:Prop
H:ConstructiveIndefiniteDescription -> C
AC_fun:FunctionalChoice
AC_depfun:DependentFunctionalChoice
A0:={A : Type & {P : A -> Prop & exists x : A, P x}}:Type
B0:=fun x : A0 => projT1 x:A0 -> Type
R0:=fun (x : A0) (y : B0 x) => projT1 (projT2 x) y:forall x : A0, B0 x -> Prop
H0:=fun x : A0 => projT2 (projT2 x):forall x : A0, (fun P : projT1 x -> Prop => exists x0 : projT1 x, P x0) (projT1 (projT2 x))
f:forall x : A0, B0 x
Hf:forall x : A0, R0 x (f x)

ConstructiveIndefiniteDescription
C:Prop
H:ConstructiveIndefiniteDescription -> C
AC_fun:FunctionalChoice
AC_depfun:DependentFunctionalChoice
A0:={A1 : Type & {P0 : A1 -> Prop & exists x : A1, P0 x}}:Type
B0:=fun x : A0 => projT1 x:A0 -> Type
R0:=fun (x : A0) (y : B0 x) => projT1 (projT2 x) y:forall x : A0, B0 x -> Prop
H0:=fun x : A0 => projT2 (projT2 x):forall x : A0, (fun P0 : projT1 x -> Prop => exists x0 : projT1 x, P0 x0) (projT1 (projT2 x))
f:forall x : A0, B0 x
Hf:forall x : A0, R0 x (f x)
A:Type
P:A -> Prop
H':exists x : A, P x

{x : A | P x}
C:Prop
H:ConstructiveIndefiniteDescription -> C
AC_fun:FunctionalChoice
AC_depfun:DependentFunctionalChoice
A0:={A1 : Type & {P0 : A1 -> Prop & exists x : A1, P0 x}}:Type
B0:=fun x : A0 => projT1 x:A0 -> Type
R0:=fun (x : A0) (y : B0 x) => projT1 (projT2 x) y:forall x : A0, B0 x -> Prop
H0:=fun x : A0 => projT2 (projT2 x):forall x : A0, (fun P0 : projT1 x -> Prop => exists x0 : projT1 x, P0 x0) (projT1 (projT2 x))
f:forall x : A0, B0 x
Hf:forall x : A0, R0 x (f x)
A:Type
P:A -> Prop
H':exists x : A, P x

P (f (existT (fun A1 : Type => {P0 : A1 -> Prop & exists x : A1, P0 x}) A (existT (fun P0 : A -> Prop => exists x : A, P0 x) P H')))
C:Prop
H:ConstructiveIndefiniteDescription -> C
AC_fun:FunctionalChoice
AC_depfun:DependentFunctionalChoice
A0:={A1 : Type & {P0 : A1 -> Prop & exists x : A1, P0 x}}:Type
B0:=fun x : A0 => projT1 x:A0 -> Type
R0:=fun (x : A0) (y : B0 x) => projT1 (projT2 x) y:forall x : A0, B0 x -> Prop
H0:=fun x : A0 => projT2 (projT2 x):forall x : A0, (fun P0 : projT1 x -> Prop => exists x0 : projT1 x, P0 x0) (projT1 (projT2 x))
f:forall x : A0, B0 x
Hf:forall x : A0, R0 x (f x)
A:Type
P:A -> Prop
H':exists x : A, P x
Hf':=Hf (existT (fun A1 : Type => {P0 : A1 -> Prop & exists x : A1, P0 x}) A (existT (fun P0 : A -> Prop => exists x : A, P0 x) P H')):R0 (existT (fun A1 : Type => {P0 : A1 -> Prop & exists x : A1, P0 x}) A (existT (fun P0 : A -> Prop => exists x : A, P0 x) P H')) (f (existT (fun A1 : Type => {P0 : A1 -> Prop & exists x : A1, P0 x}) A (existT (fun P0 : A -> Prop => exists x : A, P0 x) P H')))

P (f (existT (fun A1 : Type => {P0 : A1 -> Prop & exists x : A1, P0 x}) A (existT (fun P0 : A -> Prop => exists x : A, P0 x) P H')))
assumption. Qed.

ConstructiveIndefiniteDescription -> FunctionalChoice

ConstructiveIndefiniteDescription -> FunctionalChoice
IndefDescr:ConstructiveIndefiniteDescription
A, B:Type
R:A -> B -> Prop
H:forall x : A, exists y : B, R x y

exists f : A -> B, forall x : A, R x (f x)
IndefDescr:ConstructiveIndefiniteDescription
A, B:Type
R:A -> B -> Prop
H:forall x : A, exists y : B, R x y

forall x : A, R x (proj1_sig (IndefDescr B (R x) (H x)))
IndefDescr:ConstructiveIndefiniteDescription
A, B:Type
R:A -> B -> Prop
H:forall x0 : A, exists y : B, R x0 y
x:A

R x (proj1_sig (IndefDescr B (R x) (H x)))
apply (proj2_sig (IndefDescr B (R x) (H x))). Qed.

Non contradiction of definite description


forall C : Prop, (ConstructiveDefiniteDescription -> C) -> FunctionalRelReification -> C

forall C : Prop, (ConstructiveDefiniteDescription -> C) -> FunctionalRelReification -> C
C:Prop
H:ConstructiveDefiniteDescription -> C
FunReify:FunctionalRelReification

C
C:Prop
H:ConstructiveDefiniteDescription -> C
FunReify:FunctionalRelReification
DepFunReify:DependentFunctionalRelReification

C
C:Prop
H:ConstructiveDefiniteDescription -> C
FunReify:FunctionalRelReification
DepFunReify:DependentFunctionalRelReification
A0:={A : Type & {P : A -> Prop & exists ! x : A, P x}}:Type

C
C:Prop
H:ConstructiveDefiniteDescription -> C
FunReify:FunctionalRelReification
DepFunReify:DependentFunctionalRelReification
A0:={A : Type & {P : A -> Prop & exists ! x : A, P x}}:Type
B0:=fun x : A0 => projT1 x:A0 -> Type

C
C:Prop
H:ConstructiveDefiniteDescription -> C
FunReify:FunctionalRelReification
DepFunReify:DependentFunctionalRelReification
A0:={A : Type & {P : A -> Prop & exists ! x : A, P x}}:Type
B0:=fun x : A0 => projT1 x:A0 -> Type
R0:=fun (x : A0) (y : B0 x) => projT1 (projT2 x) y:forall x : A0, B0 x -> Prop

C
C:Prop
H:ConstructiveDefiniteDescription -> C
FunReify:FunctionalRelReification
DepFunReify:DependentFunctionalRelReification
A0:={A : Type & {P : A -> Prop & exists ! x : A, P x}}:Type
B0:=fun x : A0 => projT1 x:A0 -> Type
R0:=fun (x : A0) (y : B0 x) => projT1 (projT2 x) y:forall x : A0, B0 x -> Prop
H0:=fun x : A0 => projT2 (projT2 x):forall x : A0, (fun P : projT1 x -> Prop => exists ! x0 : projT1 x, P x0) (projT1 (projT2 x))

C
C:Prop
H:ConstructiveDefiniteDescription -> C
FunReify:FunctionalRelReification
DepFunReify:DependentFunctionalRelReification
A0:={A : Type & {P : A -> Prop & exists ! x : A, P x}}:Type
B0:=fun x : A0 => projT1 x:A0 -> Type
R0:=fun (x : A0) (y : B0 x) => projT1 (projT2 x) y:forall x : A0, B0 x -> Prop
H0:=fun x : A0 => projT2 (projT2 x):forall x : A0, (fun P : projT1 x -> Prop => exists ! x0 : projT1 x, P x0) (projT1 (projT2 x))
f:forall x : A0, B0 x
Hf:forall x : A0, R0 x (f x)

C
C:Prop
H:ConstructiveDefiniteDescription -> C
FunReify:FunctionalRelReification
DepFunReify:DependentFunctionalRelReification
A0:={A : Type & {P : A -> Prop & exists ! x : A, P x}}:Type
B0:=fun x : A0 => projT1 x:A0 -> Type
R0:=fun (x : A0) (y : B0 x) => projT1 (projT2 x) y:forall x : A0, B0 x -> Prop
H0:=fun x : A0 => projT2 (projT2 x):forall x : A0, (fun P : projT1 x -> Prop => exists ! x0 : projT1 x, P x0) (projT1 (projT2 x))
f:forall x : A0, B0 x
Hf:forall x : A0, R0 x (f x)

ConstructiveDefiniteDescription
C:Prop
H:ConstructiveDefiniteDescription -> C
FunReify:FunctionalRelReification
DepFunReify:DependentFunctionalRelReification
A0:={A1 : Type & {P0 : A1 -> Prop & exists ! x : A1, P0 x}}:Type
B0:=fun x : A0 => projT1 x:A0 -> Type
R0:=fun (x : A0) (y : B0 x) => projT1 (projT2 x) y:forall x : A0, B0 x -> Prop
H0:=fun x : A0 => projT2 (projT2 x):forall x : A0, (fun P0 : projT1 x -> Prop => exists ! x0 : projT1 x, P0 x0) (projT1 (projT2 x))
f:forall x : A0, B0 x
Hf:forall x : A0, R0 x (f x)
A:Type
P:A -> Prop
H':exists ! x : A, P x

{x : A | P x}
C:Prop
H:ConstructiveDefiniteDescription -> C
FunReify:FunctionalRelReification
DepFunReify:DependentFunctionalRelReification
A0:={A1 : Type & {P0 : A1 -> Prop & exists ! x : A1, P0 x}}:Type
B0:=fun x : A0 => projT1 x:A0 -> Type
R0:=fun (x : A0) (y : B0 x) => projT1 (projT2 x) y:forall x : A0, B0 x -> Prop
H0:=fun x : A0 => projT2 (projT2 x):forall x : A0, (fun P0 : projT1 x -> Prop => exists ! x0 : projT1 x, P0 x0) (projT1 (projT2 x))
f:forall x : A0, B0 x
Hf:forall x : A0, R0 x (f x)
A:Type
P:A -> Prop
H':exists ! x : A, P x

P (f (existT (fun A1 : Type => {P0 : A1 -> Prop & exists ! x : A1, P0 x}) A (existT (fun P0 : A -> Prop => exists ! x : A, P0 x) P H')))
C:Prop
H:ConstructiveDefiniteDescription -> C
FunReify:FunctionalRelReification
DepFunReify:DependentFunctionalRelReification
A0:={A1 : Type & {P0 : A1 -> Prop & exists ! x : A1, P0 x}}:Type
B0:=fun x : A0 => projT1 x:A0 -> Type
R0:=fun (x : A0) (y : B0 x) => projT1 (projT2 x) y:forall x : A0, B0 x -> Prop
H0:=fun x : A0 => projT2 (projT2 x):forall x : A0, (fun P0 : projT1 x -> Prop => exists ! x0 : projT1 x, P0 x0) (projT1 (projT2 x))
f:forall x : A0, B0 x
Hf:forall x : A0, R0 x (f x)
A:Type
P:A -> Prop
H':exists ! x : A, P x
Hf':=Hf (existT (fun A1 : Type => {P0 : A1 -> Prop & exists ! x : A1, P0 x}) A (existT (fun P0 : A -> Prop => exists ! x : A, P0 x) P H')):R0 (existT (fun A1 : Type => {P0 : A1 -> Prop & exists ! x : A1, P0 x}) A (existT (fun P0 : A -> Prop => exists ! x : A, P0 x) P H')) (f (existT (fun A1 : Type => {P0 : A1 -> Prop & exists ! x : A1, P0 x}) A (existT (fun P0 : A -> Prop => exists ! x : A, P0 x) P H')))

P (f (existT (fun A1 : Type => {P0 : A1 -> Prop & exists ! x : A1, P0 x}) A (existT (fun P0 : A -> Prop => exists ! x : A, P0 x) P H')))
assumption. Qed.

ConstructiveDefiniteDescription -> FunctionalRelReification

ConstructiveDefiniteDescription -> FunctionalRelReification
DefDescr:ConstructiveDefiniteDescription
A, B:Type
R:A -> B -> Prop
H:forall x : A, exists ! y : B, R x y

exists f : A -> B, forall x : A, R x (f x)
DefDescr:ConstructiveDefiniteDescription
A, B:Type
R:A -> B -> Prop
H:forall x : A, exists ! y : B, R x y

forall x : A, R x (proj1_sig (DefDescr B (R x) (H x)))
DefDescr:ConstructiveDefiniteDescription
A, B:Type
R:A -> B -> Prop
H:forall x0 : A, exists ! y : B, R x0 y
x:A

R x (proj1_sig (DefDescr B (R x) (H x)))
apply (proj2_sig (DefDescr B (R x) (H x))). Qed.
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.


ConstructiveDefiniteDescription -> (forall P : Prop, P \/ ~ P) -> forall P : Prop, {P} + {~ P}

ConstructiveDefiniteDescription -> (forall P : Prop, P \/ ~ P) -> forall P : Prop, {P} + {~ P}
Descr:ConstructiveDefiniteDescription
EM:forall P0 : Prop, P0 \/ ~ P0
P:Prop

{P} + {~ P}
Descr:ConstructiveDefiniteDescription
EM:forall P0 : Prop, P0 \/ ~ P0
P:Prop
select:=fun b : bool => if b then P else ~ P:bool -> Prop

{P} + {~ P}
Descr:ConstructiveDefiniteDescription
EM:forall P0 : Prop, P0 \/ ~ P0
P:Prop
select:=fun b : bool => if b then P else ~ P:bool -> Prop

{b : bool | select b}
Descr:ConstructiveDefiniteDescription
EM:forall P0 : Prop, P0 \/ ~ P0
P:Prop
select:=fun b : bool => if b then P else ~ P:bool -> Prop
HP:select true
{P} + {~ P}
Descr:ConstructiveDefiniteDescription
EM:forall P0 : Prop, P0 \/ ~ P0
P:Prop
select:=fun b : bool => if b then P else ~ P:bool -> Prop
HP:select false
{P} + {~ P}
Descr:forall (A : Type) (P0 : A -> Prop), (exists ! x : A, P0 x) -> {x : A | P0 x}
EM:forall P0 : Prop, P0 \/ ~ P0
P:Prop
select:=fun b : bool => if b then P else ~ P:bool -> Prop

{b : bool | select b}
Descr:ConstructiveDefiniteDescription
EM:forall P0 : Prop, P0 \/ ~ P0
P:Prop
select:=fun b : bool => if b then P else ~ P:bool -> Prop
HP:select true
{P} + {~ P}
Descr:ConstructiveDefiniteDescription
EM:forall P0 : Prop, P0 \/ ~ P0
P:Prop
select:=fun b : bool => if b then P else ~ P:bool -> Prop
HP:select false
{P} + {~ P}
Descr:forall (A : Type) (P0 : A -> Prop), (exists ! x : A, P0 x) -> {x : A | P0 x}
EM:forall P0 : Prop, P0 \/ ~ P0
P:Prop
select:=fun b : bool => if b then P else ~ P:bool -> Prop

exists ! x : bool, select x
Descr:ConstructiveDefiniteDescription
EM:forall P0 : Prop, P0 \/ ~ P0
P:Prop
select:=fun b : bool => if b then P else ~ P:bool -> Prop
HP:select true
{P} + {~ P}
Descr:ConstructiveDefiniteDescription
EM:forall P0 : Prop, P0 \/ ~ P0
P:Prop
select:=fun b : bool => if b then P else ~ P:bool -> Prop
HP:select false
{P} + {~ P}
Descr:forall (A : Type) (P0 : A -> Prop), (exists ! x : A, P0 x) -> {x : A | P0 x}
EM:forall P0 : Prop, P0 \/ ~ P0
P:Prop
select:=fun b : bool => if b then P else ~ P:bool -> Prop

exists x : bool, select x
Descr:forall (A : Type) (P0 : A -> Prop), (exists ! x : A, P0 x) -> {x : A | P0 x}
EM:forall P0 : Prop, P0 \/ ~ P0
P:Prop
select:=fun b : bool => if b then P else ~ P:bool -> Prop
uniqueness select
Descr:ConstructiveDefiniteDescription
EM:forall P0 : Prop, P0 \/ ~ P0
P:Prop
select:=fun b : bool => if b then P else ~ P:bool -> Prop
HP:select true
{P} + {~ P}
Descr:ConstructiveDefiniteDescription
EM:forall P0 : Prop, P0 \/ ~ P0
P:Prop
select:=fun b : bool => if b then P else ~ P:bool -> Prop
HP:select false
{P} + {~ P}
Descr:forall (A : Type) (P0 : A -> Prop), (exists ! x : A, P0 x) -> {x : A | P0 x}
EM:forall P0 : Prop, P0 \/ ~ P0
P:Prop
select:=fun b : bool => if b then P else ~ P:bool -> Prop
H:P

exists x : bool, select x
Descr:forall (A : Type) (P0 : A -> Prop), (exists ! x : A, P0 x) -> {x : A | P0 x}
EM:forall P0 : Prop, P0 \/ ~ P0
P:Prop
select:=fun b : bool => if b then P else ~ P:bool -> Prop
H:~ P
exists x : bool, select x
Descr:forall (A : Type) (P0 : A -> Prop), (exists ! x : A, P0 x) -> {x : A | P0 x}
EM:forall P0 : Prop, P0 \/ ~ P0
P:Prop
select:=fun b : bool => if b then P else ~ P:bool -> Prop
uniqueness select
Descr:ConstructiveDefiniteDescription
EM:forall P0 : Prop, P0 \/ ~ P0
P:Prop
select:=fun b : bool => if b then P else ~ P:bool -> Prop
HP:select true
{P} + {~ P}
Descr:ConstructiveDefiniteDescription
EM:forall P0 : Prop, P0 \/ ~ P0
P:Prop
select:=fun b : bool => if b then P else ~ P:bool -> Prop
HP:select false
{P} + {~ P}
Descr:forall (A : Type) (P0 : A -> Prop), (exists ! x : A, P0 x) -> {x : A | P0 x}
EM:forall P0 : Prop, P0 \/ ~ P0
P:Prop
select:=fun b : bool => if b then P else ~ P:bool -> Prop
H:~ P

exists x : bool, select x
Descr:forall (A : Type) (P0 : A -> Prop), (exists ! x : A, P0 x) -> {x : A | P0 x}
EM:forall P0 : Prop, P0 \/ ~ P0
P:Prop
select:=fun b : bool => if b then P else ~ P:bool -> Prop
uniqueness select
Descr:ConstructiveDefiniteDescription
EM:forall P0 : Prop, P0 \/ ~ P0
P:Prop
select:=fun b : bool => if b then P else ~ P:bool -> Prop
HP:select true
{P} + {~ P}
Descr:ConstructiveDefiniteDescription
EM:forall P0 : Prop, P0 \/ ~ P0
P:Prop
select:=fun b : bool => if b then P else ~ P:bool -> Prop
HP:select false
{P} + {~ P}
Descr:forall (A : Type) (P0 : A -> Prop), (exists ! x : A, P0 x) -> {x : A | P0 x}
EM:forall P0 : Prop, P0 \/ ~ P0
P:Prop
select:=fun b : bool => if b then P else ~ P:bool -> Prop

uniqueness select
Descr:ConstructiveDefiniteDescription
EM:forall P0 : Prop, P0 \/ ~ P0
P:Prop
select:=fun b : bool => if b then P else ~ P:bool -> Prop
HP:select true
{P} + {~ P}
Descr:ConstructiveDefiniteDescription
EM:forall P0 : Prop, P0 \/ ~ P0
P:Prop
select:=fun b : bool => if b then P else ~ P:bool -> Prop
HP:select false
{P} + {~ P}
Descr:ConstructiveDefiniteDescription
EM:forall P0 : Prop, P0 \/ ~ P0
P:Prop
select:=fun b : bool => if b then P else ~ P:bool -> Prop
HP:select true

{P} + {~ P}
Descr:ConstructiveDefiniteDescription
EM:forall P0 : Prop, P0 \/ ~ P0
P:Prop
select:=fun b : bool => if b then P else ~ P:bool -> Prop
HP:select false
{P} + {~ P}
Descr:ConstructiveDefiniteDescription
EM:forall P0 : Prop, P0 \/ ~ P0
P:Prop
select:=fun b : bool => if b then P else ~ P:bool -> Prop
HP:select false

{P} + {~ P}
right; trivial. Qed.

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

FunctionalRelReification -> (forall P : Prop, P \/ ~ P) -> forall C : Prop, ((forall P : Prop, {P} + {~ P}) -> C) -> C
FunReify:FunctionalRelReification
EM:forall P : Prop, P \/ ~ P
C:Prop
H:(forall P : Prop, {P} + {~ P}) -> C

C
intuition auto using constructive_definite_descr_excluded_middle, (relative_non_contradiction_of_definite_descr (C:=C)). Qed. (**********************************************************************)

Choice => Dependent choice => Countable choice

(* The implications below are standard *)

Require Import Arith.


FunctionalChoice -> FunctionalDependentChoice

FunctionalChoice -> FunctionalDependentChoice
FunChoice:FunctionalChoice
A:Type
R:A -> A -> Prop
HRfun:forall x : A, exists y : A, R x y
x0:A

exists f : nat -> A, f 0 = x0 /\ (forall n : nat, R (f n) (f (S n)))
FunChoice:FunctionalChoice
A:Type
R:A -> A -> Prop
g:A -> A
Rg:forall x : A, R x (g x)
x0:A

exists f : nat -> A, f 0 = x0 /\ (forall n : nat, R (f n) (f (S n)))
FunChoice:FunctionalChoice
A:Type
R:A -> A -> Prop
g:A -> A
Rg:forall x : A, R x (g x)
x0:A
f:=fix f0 (n : nat) : A := match n with | 0 => x0 | S n' => g (f0 n') end:nat -> A

exists f0 : nat -> A, f0 0 = x0 /\ (forall n : nat, R (f0 n) (f0 (S n)))
exists f; firstorder. Qed.

FunctionalDependentChoice -> FunctionalCountableChoice

FunctionalDependentChoice -> FunctionalCountableChoice
H:FunctionalDependentChoice
A:Type
R:nat -> A -> Prop
H0:forall n : nat, exists y : A, R n y

exists f : nat -> A, forall n : nat, R n (f n)
H:FunctionalDependentChoice
A:Type
R:nat -> A -> Prop
H0:forall n : nat, exists y : A, R n y
R':=fun p q : nat * A => fst q = S (fst p) /\ R (fst p) (snd q):nat * A -> nat * A -> Prop

exists f : nat -> A, forall n : nat, R n (f n)
H:FunctionalDependentChoice
A:Type
R:nat -> A -> Prop
H0:forall n : nat, exists y : A, R n y
R':=fun p q : nat * A => fst q = S (fst p) /\ R (fst p) (snd q):nat * A -> nat * A -> Prop
y0:A
Hy0:R 0 y0

exists f : nat -> A, forall n : nat, R n (f n)
H:FunctionalDependentChoice
A:Type
R:nat -> A -> Prop
H0:forall n : nat, exists y : A, R n y
R':=fun p q : nat * A => fst q = S (fst p) /\ R (fst p) (snd q):nat * A -> nat * A -> Prop
y0:A
Hy0:R 0 y0

forall x : nat * A, exists y : nat * A, R' x y
H:FunctionalDependentChoice
A:Type
R:nat -> A -> Prop
H0:forall n : nat, exists y : A, R n y
R':=fun p q : nat * A => fst q = S (fst p) /\ R (fst p) (snd q):nat * A -> nat * A -> Prop
y0:A
Hy0:R 0 y0
f:nat -> nat * A
Hf0:f 0 = (0, y0)
HfS:forall n : nat, R' (f n) (f (S n))
exists f0 : nat -> A, forall n : nat, R n (f0 n)
H:FunctionalDependentChoice
A:Type
R:nat -> A -> Prop
H0:forall n : nat, exists y1 : A, R n y1
R':=fun p q : nat * A => fst q = S (fst p) /\ R (fst p) (snd q):nat * A -> nat * A -> Prop
y0:A
Hy0:R 0 y0
x:(nat * A)%type
y:A
Hy:R (fst x) y

exists y1 : nat * A, R' x y1
H:FunctionalDependentChoice
A:Type
R:nat -> A -> Prop
H0:forall n : nat, exists y : A, R n y
R':=fun p q : nat * A => fst q = S (fst p) /\ R (fst p) (snd q):nat * A -> nat * A -> Prop
y0:A
Hy0:R 0 y0
f:nat -> nat * A
Hf0:f 0 = (0, y0)
HfS:forall n : nat, R' (f n) (f (S n))
exists f0 : nat -> A, forall n : nat, R n (f0 n)
H:FunctionalDependentChoice
A:Type
R:nat -> A -> Prop
H0:forall n : nat, exists y1 : A, R n y1
R':=fun p q : nat * A => fst q = S (fst p) /\ R (fst p) (snd q):nat * A -> nat * A -> Prop
y0:A
Hy0:R 0 y0
x:(nat * A)%type
y:A
Hy:R (fst x) y

R' x (S (fst x), y)
H:FunctionalDependentChoice
A:Type
R:nat -> A -> Prop
H0:forall n : nat, exists y : A, R n y
R':=fun p q : nat * A => fst q = S (fst p) /\ R (fst p) (snd q):nat * A -> nat * A -> Prop
y0:A
Hy0:R 0 y0
f:nat -> nat * A
Hf0:f 0 = (0, y0)
HfS:forall n : nat, R' (f n) (f (S n))
exists f0 : nat -> A, forall n : nat, R n (f0 n)
H:FunctionalDependentChoice
A:Type
R:nat -> A -> Prop
H0:forall n : nat, exists y1 : A, R n y1
R':=fun p q : nat * A => fst q = S (fst p) /\ R (fst p) (snd q):nat * A -> nat * A -> Prop
y0:A
Hy0:R 0 y0
x:(nat * A)%type
y:A
Hy:R (fst x) y

fst (S (fst x), y) = S (fst x) /\ R (fst x) (snd (S (fst x), y))
H:FunctionalDependentChoice
A:Type
R:nat -> A -> Prop
H0:forall n : nat, exists y : A, R n y
R':=fun p q : nat * A => fst q = S (fst p) /\ R (fst p) (snd q):nat * A -> nat * A -> Prop
y0:A
Hy0:R 0 y0
f:nat -> nat * A
Hf0:f 0 = (0, y0)
HfS:forall n : nat, R' (f n) (f (S n))
exists f0 : nat -> A, forall n : nat, R n (f0 n)
H:FunctionalDependentChoice
A:Type
R:nat -> A -> Prop
H0:forall n : nat, exists y : A, R n y
R':=fun p q : nat * A => fst q = S (fst p) /\ R (fst p) (snd q):nat * A -> nat * A -> Prop
y0:A
Hy0:R 0 y0
f:nat -> nat * A
Hf0:f 0 = (0, y0)
HfS:forall n : nat, R' (f n) (f (S n))

exists f0 : nat -> A, forall n : nat, R n (f0 n)
H:FunctionalDependentChoice
A:Type
R:nat -> A -> Prop
H0:forall n : nat, exists y : A, R n y
R':=fun p q : nat * A => fst q = S (fst p) /\ R (fst p) (snd q):nat * A -> nat * A -> Prop
y0:A
Hy0:R 0 y0
f:nat -> nat * A
Hf0:f 0 = (0, y0)
HfS:forall n : nat, R' (f n) (f (S n))

forall n : nat, fst (f n) = n
H:FunctionalDependentChoice
A:Type
R:nat -> A -> Prop
H0:forall n : nat, exists y : A, R n y
R':=fun p q : nat * A => fst q = S (fst p) /\ R (fst p) (snd q):nat * A -> nat * A -> Prop
y0:A
Hy0:R 0 y0
f:nat -> nat * A
Hf0:f 0 = (0, y0)
HfS:forall n : nat, R' (f n) (f (S n))
Heq:forall n : nat, fst (f n) = n
exists f0 : nat -> A, forall n : nat, R n (f0 n)
H:FunctionalDependentChoice
A:Type
R:nat -> A -> Prop
H0:forall n : nat, exists y : A, R n y
R':=fun p q : nat * A => fst q = S (fst p) /\ R (fst p) (snd q):nat * A -> nat * A -> Prop
y0:A
Hy0:R 0 y0
f:nat -> nat * A
Hf0:f 0 = (0, y0)
HfS:forall n : nat, R' (f n) (f (S n))

fst (f 0) = 0
H:FunctionalDependentChoice
A:Type
R:nat -> A -> Prop
H0:forall n0 : nat, exists y : A, R n0 y
R':=fun p q : nat * A => fst q = S (fst p) /\ R (fst p) (snd q):nat * A -> nat * A -> Prop
y0:A
Hy0:R 0 y0
f:nat -> nat * A
Hf0:f 0 = (0, y0)
HfS:forall n0 : nat, R' (f n0) (f (S n0))
n:nat
IHn:fst (f n) = n
fst (f (S n)) = S n
H:FunctionalDependentChoice
A:Type
R:nat -> A -> Prop
H0:forall n : nat, exists y : A, R n y
R':=fun p q : nat * A => fst q = S (fst p) /\ R (fst p) (snd q):nat * A -> nat * A -> Prop
y0:A
Hy0:R 0 y0
f:nat -> nat * A
Hf0:f 0 = (0, y0)
HfS:forall n : nat, R' (f n) (f (S n))
Heq:forall n : nat, fst (f n) = n
exists f0 : nat -> A, forall n : nat, R n (f0 n)
H:FunctionalDependentChoice
A:Type
R:nat -> A -> Prop
H0:forall n0 : nat, exists y : A, R n0 y
R':=fun p q : nat * A => fst q = S (fst p) /\ R (fst p) (snd q):nat * A -> nat * A -> Prop
y0:A
Hy0:R 0 y0
f:nat -> nat * A
Hf0:f 0 = (0, y0)
HfS:forall n0 : nat, R' (f n0) (f (S n0))
n:nat
IHn:fst (f n) = n

fst (f (S n)) = S n
H:FunctionalDependentChoice
A:Type
R:nat -> A -> Prop
H0:forall n : nat, exists y : A, R n y
R':=fun p q : nat * A => fst q = S (fst p) /\ R (fst p) (snd q):nat * A -> nat * A -> Prop
y0:A
Hy0:R 0 y0
f:nat -> nat * A
Hf0:f 0 = (0, y0)
HfS:forall n : nat, R' (f n) (f (S n))
Heq:forall n : nat, fst (f n) = n
exists f0 : nat -> A, forall n : nat, R n (f0 n)
H:FunctionalDependentChoice
A:Type
R:nat -> A -> Prop
H0:forall n : nat, exists y : A, R n y
R':=fun p q : nat * A => fst q = S (fst p) /\ R (fst p) (snd q):nat * A -> nat * A -> Prop
y0:A
Hy0:R 0 y0
f:nat -> nat * A
Hf0:f 0 = (0, y0)
HfS:forall n : nat, R' (f n) (f (S n))
Heq:forall n : nat, fst (f n) = n

exists f0 : nat -> A, forall n : nat, R n (f0 n)
H:FunctionalDependentChoice
A:Type
R:nat -> A -> Prop
H0:forall n : nat, exists y : A, R n y
R':=fun p q : nat * A => fst q = S (fst p) /\ R (fst p) (snd q):nat * A -> nat * A -> Prop
y0:A
Hy0:R 0 y0
f:nat -> nat * A
Hf0:f 0 = (0, y0)
HfS:forall n : nat, R' (f n) (f (S n))
Heq:forall n : nat, fst (f n) = n

forall n : nat, R n (snd (f (S n)))
H:FunctionalDependentChoice
A:Type
R:nat -> A -> Prop
H0:forall n : nat, exists y : A, R n y
R':=fun p q : nat * A => fst q = S (fst p) /\ R (fst p) (snd q):nat * A -> nat * A -> Prop
y0:A
Hy0:R 0 y0
f:nat -> nat * A
Hf0:f 0 = (0, y0)
HfS:forall n : nat, R' (f n) (f (S n))
Heq:forall n : nat, fst (f n) = n
n':nat

R n' (snd (f (S n')))
H:FunctionalDependentChoice
A:Type
R:nat -> A -> Prop
H0:forall n : nat, exists y : A, R n y
R':=fun p q : nat * A => fst q = S (fst p) /\ R (fst p) (snd q):nat * A -> nat * A -> Prop
y0:A
Hy0:R 0 y0
f:nat -> nat * A
Hf0:f 0 = (0, y0)
n':nat
HfS:R' (f n') (f (S n'))
Heq:forall n : nat, fst (f n) = n

R n' (snd (f (S n')))
H:FunctionalDependentChoice
A:Type
R:nat -> A -> Prop
H0:forall n : nat, exists y : A, R n y
R':=fun p q : nat * A => fst q = S (fst p) /\ R (fst p) (snd q):nat * A -> nat * A -> Prop
y0:A
Hy0:R 0 y0
f:nat -> nat * A
Hf0:f 0 = (0, y0)
n':nat
HR:R (fst (f n')) (snd (f (S n')))
Heq:forall n : nat, fst (f n) = n

R n' (snd (f (S n')))
H:FunctionalDependentChoice
A:Type
R:nat -> A -> Prop
H0:forall n : nat, exists y : A, R n y
R':=fun p q : nat * A => fst q = S (fst p) /\ R (fst p) (snd q):nat * A -> nat * A -> Prop
y0:A
Hy0:R 0 y0
f:nat -> nat * A
Hf0:f 0 = (0, y0)
n':nat
HR:R n' (snd (f (S n')))
Heq:forall n : nat, fst (f n) = n

R n' (snd (f (S n')))
assumption. Qed. (**********************************************************************)

About the axiom of choice over setoids

Require Import ClassicalFacts PropExtensionalityFacts.

(**********************************************************************)

Consequences of the choice of a representative in an equivalence class


RepresentativeFunctionalChoice -> Type -> exists h : Prop -> Prop, forall P : Prop, (P <-> h P) /\ (forall Q : Prop, P <-> Q -> h P = h Q)

RepresentativeFunctionalChoice -> Type -> exists h : Prop -> Prop, forall P : Prop, (P <-> h P) /\ (forall Q : Prop, P <-> Q -> h P = h Q)
ReprFunChoice:RepresentativeFunctionalChoice
A:Type

exists h : Prop -> Prop, forall P : Prop, (P <-> h P) /\ (forall Q : Prop, P <-> Q -> h P = h Q)
ReprFunChoice:RepresentativeFunctionalChoice
A:Type
R:=fun P Q : Prop => P <-> Q:Prop -> Prop -> Prop

exists h : Prop -> Prop, forall P : Prop, (P <-> h P) /\ (forall Q : Prop, P <-> Q -> h P = h Q)
ReprFunChoice:RepresentativeFunctionalChoice
A:Type
R:=fun P Q : Prop => P <-> Q:Prop -> Prop -> Prop
Hequiv:Equivalence R

exists h : Prop -> Prop, forall P : Prop, (P <-> h P) /\ (forall Q : Prop, P <-> Q -> h P = h Q)
apply (ReprFunChoice _ R Hequiv). Qed.

RepresentativeFunctionalChoice -> ExtensionalPredicateRepresentative

RepresentativeFunctionalChoice -> ExtensionalPredicateRepresentative
ReprFunChoice:RepresentativeFunctionalChoice
A:Type

exists h : (A -> Prop) -> A -> Prop, forall P : A -> Prop, (forall x : A, P x <-> h P x) /\ (forall Q : A -> Prop, (forall x : A, P x <-> Q x) -> h P = h Q)
ReprFunChoice:RepresentativeFunctionalChoice
A:Type
R:=fun P Q : A -> Prop => forall x : A, P x <-> Q x:(A -> Prop) -> (A -> Prop) -> Prop

exists h : (A -> Prop) -> A -> Prop, forall P : A -> Prop, (forall x : A, P x <-> h P x) /\ (forall Q : A -> Prop, (forall x : A, P x <-> Q x) -> h P = h Q)
ReprFunChoice:RepresentativeFunctionalChoice
A:Type
R:=fun P Q : A -> Prop => forall x : A, P x <-> Q x:(A -> Prop) -> (A -> Prop) -> Prop
Hequiv:Equivalence R

exists h : (A -> Prop) -> A -> Prop, forall P : A -> Prop, (forall x : A, P x <-> h P x) /\ (forall Q : A -> Prop, (forall x : A, P x <-> Q x) -> h P = h Q)
apply (ReprFunChoice _ R Hequiv). Qed.

RepresentativeFunctionalChoice -> ExtensionalFunctionRepresentative

RepresentativeFunctionalChoice -> ExtensionalFunctionRepresentative
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type

exists h : (A -> B) -> A -> B, forall f : A -> B, (forall x : A, f x = h f x) /\ (forall g : A -> B, (forall x : A, f x = g x) -> h f = h g)
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
R:=fun f g : A -> B => forall x : A, f x = g x:(A -> B) -> (A -> B) -> Prop

exists h : (A -> B) -> A -> B, forall f : A -> B, (forall x : A, f x = h f x) /\ (forall g : A -> B, (forall x : A, f x = g x) -> h f = h g)
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
R:=fun f g : A -> B => forall x : A, f x = g x:(A -> B) -> (A -> B) -> Prop

Equivalence R
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
R:=fun f g : A -> B => forall x : A, f x = g x:(A -> B) -> (A -> B) -> Prop
Hequiv:Equivalence R
exists h : (A -> B) -> A -> B, forall f : A -> B, (forall x : A, f x = h f x) /\ (forall g : A -> B, (forall x : A, f x = g x) -> h f = h g)
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
R:=fun f g : A -> B => forall x : A, f x = g x:(A -> B) -> (A -> B) -> Prop

Equivalence R
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
R:=fun f g : A -> B => forall x : A, f x = g x:(A -> B) -> (A -> B) -> Prop

Transitive R
firstorder using eq_trans.
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
R:=fun f g : A -> B => forall x : A, f x = g x:(A -> B) -> (A -> B) -> Prop
Hequiv:Equivalence R

exists h : (A -> B) -> A -> B, forall f : A -> B, (forall x : A, f x = h f x) /\ (forall g : A -> B, (forall x : A, f x = g x) -> h f = h g)
apply (ReprFunChoice _ R Hequiv). Qed.

This is a variant of Diaconescu and Goodman-Myhill theorems


RepresentativeFunctionalChoice -> ExcludedMiddle

RepresentativeFunctionalChoice -> ExcludedMiddle
ReprFunChoice:RepresentativeFunctionalChoice

ExcludedMiddle
apply representative_boolean_partition_imp_excluded_middle, ReprFunChoice. Qed.

RepresentativeFunctionalChoice -> RelationalChoice

RepresentativeFunctionalChoice -> RelationalChoice
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x : A, exists y : B, T x y

exists R' : A -> B -> Prop, subrelation R' T /\ (forall x : A, exists ! y : B, R' x y)
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x : A, exists y : B, T x y
D:=(A * B)%type:Type

exists R' : A -> B -> Prop, subrelation R' T /\ (forall x : A, exists ! y : B, R' x y)
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x : A, exists y : B, T x y
D:=(A * B)%type:Type
R:=fun z z' : D => let x := fst z in let x' := fst z' in let y := snd z in let y' := snd z' in x = x' /\ (T x y -> y = y' \/ T x y') /\ (T x y' -> y = y' \/ T x y):D -> D -> Prop

exists R' : A -> B -> Prop, subrelation R' T /\ (forall x : A, exists ! y : B, R' x y)
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x : A, exists y : B, T x y
D:=(A * B)%type:Type
R:=fun z z' : D => let x := fst z in let x' := fst z' in let y := snd z in let y' := snd z' in x = x' /\ (T x y -> y = y' \/ T x y') /\ (T x y' -> y = y' \/ T x y):D -> D -> Prop

Equivalence R
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x : A, exists y : B, T x y
D:=(A * B)%type:Type
R:=fun z z' : D => let x := fst z in let x' := fst z' in let y := snd z in let y' := snd z' in x = x' /\ (T x y -> y = y' \/ T x y') /\ (T x y' -> y = y' \/ T x y):D -> D -> Prop
Hequiv:Equivalence R
exists R' : A -> B -> Prop, subrelation R' T /\ (forall x : A, exists ! y : B, R' x y)
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x : A, exists y : B, T x y
D:=(A * B)%type:Type
R:=fun z z' : D => let x := fst z in let x' := fst z' in let y := snd z in let y' := snd z' in x = x' /\ (T x y -> y = y' \/ T x y') /\ (T x y' -> y = y' \/ T x y):D -> D -> Prop

Equivalence R
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x : A, exists y : B, T x y
D:=(A * B)%type:Type
R:=fun z z' : D => let x := fst z in let x' := fst z' in let y := snd z in let y' := snd z' in x = x' /\ (T x y -> y = y' \/ T x y') /\ (T x y' -> y = y' \/ T x y):D -> D -> Prop

Reflexive R
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x : A, exists y : B, T x y
D:=(A * B)%type:Type
R:=fun z z' : D => let x := fst z in let x' := fst z' in let y := snd z in let y' := snd z' in x = x' /\ (T x y -> y = y' \/ T x y') /\ (T x y' -> y = y' \/ T x y):D -> D -> Prop
Symmetric R
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x : A, exists y : B, T x y
D:=(A * B)%type:Type
R:=fun z z' : D => let x := fst z in let x' := fst z' in let y := snd z in let y' := snd z' in x = x' /\ (T x y -> y = y' \/ T x y') /\ (T x y' -> y = y' \/ T x y):D -> D -> Prop
Transitive R
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x : A, exists y : B, T x y
D:=(A * B)%type:Type
R:=fun z z' : D => let x := fst z in let x' := fst z' in let y := snd z in let y' := snd z' in x = x' /\ (T x y -> y = y' \/ T x y') /\ (T x y' -> y = y' \/ T x y):D -> D -> Prop

Reflexive R
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y : B, T x0 y
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y := snd z in let y' := snd z' in x0 = x' /\ (T x0 y -> y = y' \/ T x0 y') /\ (T x0 y' -> y = y' \/ T x0 y):D -> D -> Prop
x:D

fst x = fst x
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y : B, T x0 y
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y := snd z in let y' := snd z' in x0 = x' /\ (T x0 y -> y = y' \/ T x0 y') /\ (T x0 y' -> y = y' \/ T x0 y):D -> D -> Prop
x:D
(T (fst x) (snd x) -> snd x = snd x \/ T (fst x) (snd x)) /\ (T (fst x) (snd x) -> snd x = snd x \/ T (fst x) (snd x))
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y : B, T x0 y
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y := snd z in let y' := snd z' in x0 = x' /\ (T x0 y -> y = y' \/ T x0 y') /\ (T x0 y' -> y = y' \/ T x0 y):D -> D -> Prop
x:D

(T (fst x) (snd x) -> snd x = snd x \/ T (fst x) (snd x)) /\ (T (fst x) (snd x) -> snd x = snd x \/ T (fst x) (snd x))
firstorder.
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x : A, exists y : B, T x y
D:=(A * B)%type:Type
R:=fun z z' : D => let x := fst z in let x' := fst z' in let y := snd z in let y' := snd z' in x = x' /\ (T x y -> y = y' \/ T x y') /\ (T x y' -> y = y' \/ T x y):D -> D -> Prop

Symmetric R
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x'0 := fst z' in let y0 := snd z in let y'0 := snd z' in x0 = x'0 /\ (T x0 y0 -> y0 = y'0 \/ T x0 y'0) /\ (T x0 y'0 -> y0 = y'0 \/ T x0 y0):D -> D -> Prop
x:A
y:B
x':A
y':B
H1:fst (x, y) = fst (x', y')
H2:T (fst (x, y)) (snd (x, y)) -> snd (x, y) = snd (x', y') \/ T (fst (x, y)) (snd (x', y'))
H2':T (fst (x, y)) (snd (x', y')) -> snd (x, y) = snd (x', y') \/ T (fst (x, y)) (snd (x, y))

R (x', y') (x, y)
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x'0 := fst z' in let y0 := snd z in let y'0 := snd z' in x0 = x'0 /\ (T x0 y0 -> y0 = y'0 \/ T x0 y'0) /\ (T x0 y'0 -> y0 = y'0 \/ T x0 y0):D -> D -> Prop
x:A
y:B
x':A
y':B
H1:fst (x, y) = fst (x', y')
H2:T (fst (x, y)) (snd (x, y)) -> snd (x, y) = snd (x', y') \/ T (fst (x, y)) (snd (x', y'))
H2':T (fst (x, y)) (snd (x', y')) -> snd (x, y) = snd (x', y') \/ T (fst (x, y)) (snd (x, y))

fst (x', y') = fst (x, y)
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x'0 := fst z' in let y0 := snd z in let y'0 := snd z' in x0 = x'0 /\ (T x0 y0 -> y0 = y'0 \/ T x0 y'0) /\ (T x0 y'0 -> y0 = y'0 \/ T x0 y0):D -> D -> Prop
x:A
y:B
x':A
y':B
H1:fst (x, y) = fst (x', y')
H2:T (fst (x, y)) (snd (x, y)) -> snd (x, y) = snd (x', y') \/ T (fst (x, y)) (snd (x', y'))
H2':T (fst (x, y)) (snd (x', y')) -> snd (x, y) = snd (x', y') \/ T (fst (x, y)) (snd (x, y))
(T (fst (x', y')) (snd (x', y')) -> snd (x', y') = snd (x, y) \/ T (fst (x', y')) (snd (x, y))) /\ (T (fst (x', y')) (snd (x, y)) -> snd (x', y') = snd (x, y) \/ T (fst (x', y')) (snd (x', y')))
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x'0 := fst z' in let y0 := snd z in let y'0 := snd z' in x0 = x'0 /\ (T x0 y0 -> y0 = y'0 \/ T x0 y'0) /\ (T x0 y'0 -> y0 = y'0 \/ T x0 y0):D -> D -> Prop
x:A
y:B
x':A
y':B
H1:fst (x, y) = fst (x', y')
H2:T (fst (x, y)) (snd (x, y)) -> snd (x, y) = snd (x', y') \/ T (fst (x, y)) (snd (x', y'))
H2':T (fst (x, y)) (snd (x', y')) -> snd (x, y) = snd (x', y') \/ T (fst (x, y)) (snd (x, y))

(T (fst (x', y')) (snd (x', y')) -> snd (x', y') = snd (x, y) \/ T (fst (x', y')) (snd (x, y))) /\ (T (fst (x', y')) (snd (x, y)) -> snd (x', y') = snd (x, y) \/ T (fst (x', y')) (snd (x', y')))
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x'0 := fst z' in let y0 := snd z in let y'0 := snd z' in x0 = x'0 /\ (T x0 y0 -> y0 = y'0 \/ T x0 y'0) /\ (T x0 y'0 -> y0 = y'0 \/ T x0 y0):D -> D -> Prop
x:A
y:B
x':A
y':B
H1:x = x'
H2:T x (snd (x, y)) -> snd (x, y) = snd (x', y') \/ T x (snd (x', y'))
H2':T x (snd (x', y')) -> snd (x, y) = snd (x', y') \/ T x (snd (x, y))

(T x' (snd (x', y')) -> snd (x', y') = snd (x, y) \/ T x' (snd (x, y))) /\ (T x' (snd (x, y)) -> snd (x', y') = snd (x, y) \/ T x' (snd (x', y')))
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x'0 := fst z' in let y0 := snd z in let y'0 := snd z' in x0 = x'0 /\ (T x0 y0 -> y0 = y'0 \/ T x0 y'0) /\ (T x0 y'0 -> y0 = y'0 \/ T x0 y0):D -> D -> Prop
x:A
y:B
x':A
y':B
H1:x = x'
H2:T x y -> y = y' \/ T x y'
H2':T x y' -> y = y' \/ T x y

(T x' y' -> y' = y \/ T x' y) /\ (T x' y -> y' = y \/ T x' y')
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y0 := snd z in let y'0 := snd z' in x0 = x' /\ (T x0 y0 -> y0 = y'0 \/ T x0 y'0) /\ (T x0 y'0 -> y0 = y'0 \/ T x0 y0):D -> D -> Prop
x:A
y, y':B
H2:T x y -> y = y' \/ T x y'
H2':T x y' -> y = y' \/ T x y

(T x y' -> y' = y \/ T x y) /\ (T x y -> y' = y \/ T x y')
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y0 := snd z in let y'0 := snd z' in x0 = x' /\ (T x0 y0 -> y0 = y'0 \/ T x0 y'0) /\ (T x0 y'0 -> y0 = y'0 \/ T x0 y0):D -> D -> Prop
x:A
y, y':B
H2:T x y -> y = y' \/ T x y'
H2':T x y' -> y = y' \/ T x y
H:T x y'

y' = y \/ T x y
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y0 := snd z in let y'0 := snd z' in x0 = x' /\ (T x0 y0 -> y0 = y'0 \/ T x0 y'0) /\ (T x0 y'0 -> y0 = y'0 \/ T x0 y0):D -> D -> Prop
x:A
y, y':B
H2:T x y -> y = y' \/ T x y'
H2':T x y' -> y = y' \/ T x y
H:T x y
y' = y \/ T x y'
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y0 := snd z in let y'0 := snd z' in x0 = x' /\ (T x0 y0 -> y0 = y'0 \/ T x0 y'0) /\ (T x0 y'0 -> y0 = y'0 \/ T x0 y0):D -> D -> Prop
x:A
y, y':B
H2:T x y -> y = y' \/ T x y'
H2':T x y' -> y = y' \/ T x y
H:T x y'

y' = y \/ T x y
destruct (H2' H); firstorder.
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y0 := snd z in let y'0 := snd z' in x0 = x' /\ (T x0 y0 -> y0 = y'0 \/ T x0 y'0) /\ (T x0 y'0 -> y0 = y'0 \/ T x0 y0):D -> D -> Prop
x:A
y, y':B
H2:T x y -> y = y' \/ T x y'
H2':T x y' -> y = y' \/ T x y
H:T x y

y' = y \/ T x y'
destruct (H2 H); firstorder.
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x : A, exists y : B, T x y
D:=(A * B)%type:Type
R:=fun z z' : D => let x := fst z in let x' := fst z' in let y := snd z in let y' := snd z' in x = x' /\ (T x y -> y = y' \/ T x y') /\ (T x y' -> y = y' \/ T x y):D -> D -> Prop

Transitive R
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x'0 := fst z' in let y0 := snd z in let y'0 := snd z' in x0 = x'0 /\ (T x0 y0 -> y0 = y'0 \/ T x0 y'0) /\ (T x0 y'0 -> y0 = y'0 \/ T x0 y0):D -> D -> Prop
x:A
y:B
x':A
y':B
x'':A
y'':B
H1:fst (x, y) = fst (x', y')
H2:T (fst (x, y)) (snd (x, y)) -> snd (x, y) = snd (x', y') \/ T (fst (x, y)) (snd (x', y'))
H2':T (fst (x, y)) (snd (x', y')) -> snd (x, y) = snd (x', y') \/ T (fst (x, y)) (snd (x, y))
H3:fst (x', y') = fst (x'', y'')
H4:T (fst (x', y')) (snd (x', y')) -> snd (x', y') = snd (x'', y'') \/ T (fst (x', y')) (snd (x'', y''))
H4':T (fst (x', y')) (snd (x'', y'')) -> snd (x', y') = snd (x'', y'') \/ T (fst (x', y')) (snd (x', y'))

R (x, y) (x'', y'')
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x'0 := fst z' in let y0 := snd z in let y'0 := snd z' in x0 = x'0 /\ (T x0 y0 -> y0 = y'0 \/ T x0 y'0) /\ (T x0 y'0 -> y0 = y'0 \/ T x0 y0):D -> D -> Prop
x:A
y:B
x':A
y':B
x'':A
y'':B
H1:x = x'
H2:T x (snd (x, y)) -> snd (x, y) = snd (x', y') \/ T x (snd (x', y'))
H2':T x (snd (x', y')) -> snd (x, y) = snd (x', y') \/ T x (snd (x, y))
H3:x' = x''
H4:T x' (snd (x', y')) -> snd (x', y') = snd (x'', y'') \/ T x' (snd (x'', y''))
H4':T x' (snd (x'', y'')) -> snd (x', y') = snd (x'', y'') \/ T x' (snd (x', y'))

R (x, y) (x'', y'')
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x'0 := fst z' in let y0 := snd z in let y'0 := snd z' in x0 = x'0 /\ (T x0 y0 -> y0 = y'0 \/ T x0 y'0) /\ (T x0 y'0 -> y0 = y'0 \/ T x0 y0):D -> D -> Prop
x:A
y:B
x':A
y':B
x'':A
y'':B
H1:x = x'
H2:T x y -> y = y' \/ T x y'
H2':T x y' -> y = y' \/ T x y
H3:x' = x''
H4:T x' y' -> y' = y'' \/ T x' y''
H4':T x' y'' -> y' = y'' \/ T x' y'

R (x, y) (x'', y'')
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y0 := snd z in let y'0 := snd z' in x0 = x' /\ (T x0 y0 -> y0 = y'0 \/ T x0 y'0) /\ (T x0 y'0 -> y0 = y'0 \/ T x0 y0):D -> D -> Prop
x:A
y, y', y'':B
H2:T x y -> y = y' \/ T x y'
H2':T x y' -> y = y' \/ T x y
H4':T x y'' -> y' = y'' \/ T x y'
H4:T x y' -> y' = y'' \/ T x y''

R (x, y) (x, y'')
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y0 := snd z in let y'0 := snd z' in x0 = x' /\ (T x0 y0 -> y0 = y'0 \/ T x0 y'0) /\ (T x0 y'0 -> y0 = y'0 \/ T x0 y0):D -> D -> Prop
x:A
y, y', y'':B
H2:T x y -> y = y' \/ T x y'
H2':T x y' -> y = y' \/ T x y
H4':T x y'' -> y' = y'' \/ T x y'
H4:T x y' -> y' = y'' \/ T x y''

fst (x, y) = fst (x, y'')
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y0 := snd z in let y'0 := snd z' in x0 = x' /\ (T x0 y0 -> y0 = y'0 \/ T x0 y'0) /\ (T x0 y'0 -> y0 = y'0 \/ T x0 y0):D -> D -> Prop
x:A
y, y', y'':B
H2:T x y -> y = y' \/ T x y'
H2':T x y' -> y = y' \/ T x y
H4':T x y'' -> y' = y'' \/ T x y'
H4:T x y' -> y' = y'' \/ T x y''
(T (fst (x, y)) (snd (x, y)) -> snd (x, y) = snd (x, y'') \/ T (fst (x, y)) (snd (x, y''))) /\ (T (fst (x, y)) (snd (x, y'')) -> snd (x, y) = snd (x, y'') \/ T (fst (x, y)) (snd (x, y)))
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y0 := snd z in let y'0 := snd z' in x0 = x' /\ (T x0 y0 -> y0 = y'0 \/ T x0 y'0) /\ (T x0 y'0 -> y0 = y'0 \/ T x0 y0):D -> D -> Prop
x:A
y, y', y'':B
H2:T x y -> y = y' \/ T x y'
H2':T x y' -> y = y' \/ T x y
H4':T x y'' -> y' = y'' \/ T x y'
H4:T x y' -> y' = y'' \/ T x y''

(T (fst (x, y)) (snd (x, y)) -> snd (x, y) = snd (x, y'') \/ T (fst (x, y)) (snd (x, y''))) /\ (T (fst (x, y)) (snd (x, y'')) -> snd (x, y) = snd (x, y'') \/ T (fst (x, y)) (snd (x, y)))
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y0 := snd z in let y'0 := snd z' in x0 = x' /\ (T x0 y0 -> y0 = y'0 \/ T x0 y'0) /\ (T x0 y'0 -> y0 = y'0 \/ T x0 y0):D -> D -> Prop
x:A
y, y', y'':B
H2:T x y -> y = y' \/ T x y'
H2':T x y' -> y = y' \/ T x y
H4':T x y'' -> y' = y'' \/ T x y'
H4:T x y' -> y' = y'' \/ T x y''
H:T (fst (x, y)) (snd (x, y))

snd (x, y) = snd (x, y'') \/ T (fst (x, y)) (snd (x, y''))
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y0 := snd z in let y'0 := snd z' in x0 = x' /\ (T x0 y0 -> y0 = y'0 \/ T x0 y'0) /\ (T x0 y'0 -> y0 = y'0 \/ T x0 y0):D -> D -> Prop
x:A
y, y', y'':B
H2:T x y -> y = y' \/ T x y'
H2':T x y' -> y = y' \/ T x y
H4':T x y'' -> y' = y'' \/ T x y'
H4:T x y' -> y' = y'' \/ T x y''
H:T (fst (x, y)) (snd (x, y''))
snd (x, y) = snd (x, y'') \/ T (fst (x, y)) (snd (x, y))
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y0 := snd z in let y'0 := snd z' in x0 = x' /\ (T x0 y0 -> y0 = y'0 \/ T x0 y'0) /\ (T x0 y'0 -> y0 = y'0 \/ T x0 y0):D -> D -> Prop
x:A
y, y', y'':B
H2:T x y -> y = y' \/ T x y'
H2':T x y' -> y = y' \/ T x y
H4':T x y'' -> y' = y'' \/ T x y'
H4:T x y' -> y' = y'' \/ T x y''
H:T (fst (x, y)) (snd (x, y))

snd (x, y) = snd (x, y'') \/ T (fst (x, y)) (snd (x, y''))
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y0 := snd z in let y'0 := snd z' in x0 = x' /\ (T x0 y0 -> y0 = y'0 \/ T x0 y'0) /\ (T x0 y'0 -> y0 = y'0 \/ T x0 y0):D -> D -> Prop
x:A
y, y', y'':B
H2:T x y -> y = y' \/ T x y'
H2':T x y' -> y = y' \/ T x y
H4':T x y'' -> y' = y'' \/ T x y'
H4:T x y' -> y' = y'' \/ T x y''
H:T x (snd (x, y))

snd (x, y) = snd (x, y'') \/ T x (snd (x, y''))
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y0 := snd z in let y'0 := snd z' in x0 = x' /\ (T x0 y0 -> y0 = y'0 \/ T x0 y'0) /\ (T x0 y'0 -> y0 = y'0 \/ T x0 y0):D -> D -> Prop
x:A
y, y', y'':B
H2:T x y -> y = y' \/ T x y'
H2':T x y' -> y = y' \/ T x y
H4':T x y'' -> y' = y'' \/ T x y'
H4:T x y' -> y' = y'' \/ T x y''
H:T x y

y = y'' \/ T x y''
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y0 := snd z in let y' := snd z' in x0 = x' /\ (T x0 y0 -> y0 = y' \/ T x0 y') /\ (T x0 y' -> y0 = y' \/ T x0 y0):D -> D -> Prop
x:A
y, y'':B
H4:T x y -> y = y'' \/ T x y''
H4':T x y'' -> y = y'' \/ T x y
H2', H2:T x y -> y = y \/ T x y
H:T x y

y = y'' \/ T x y''
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y0 := snd z in let y'0 := snd z' in x0 = x' /\ (T x0 y0 -> y0 = y'0 \/ T x0 y'0) /\ (T x0 y'0 -> y0 = y'0 \/ T x0 y0):D -> D -> Prop
x:A
y, y', y'':B
H2:T x y -> y = y' \/ T x y'
H2':T x y' -> y = y' \/ T x y
H4':T x y'' -> y' = y'' \/ T x y'
H4:T x y' -> y' = y'' \/ T x y''
H:T x y
H0:T x y'
y = y'' \/ T x y''
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y0 := snd z in let y' := snd z' in x0 = x' /\ (T x0 y0 -> y0 = y' \/ T x0 y') /\ (T x0 y' -> y0 = y' \/ T x0 y0):D -> D -> Prop
x:A
y, y'':B
H4:T x y -> y = y'' \/ T x y''
H4':T x y'' -> y = y'' \/ T x y
H2', H2:T x y -> y = y \/ T x y
H:T x y

y = y'' \/ T x y''
destruct (H4 H); firstorder.
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y0 := snd z in let y'0 := snd z' in x0 = x' /\ (T x0 y0 -> y0 = y'0 \/ T x0 y'0) /\ (T x0 y'0 -> y0 = y'0 \/ T x0 y0):D -> D -> Prop
x:A
y, y', y'':B
H2:T x y -> y = y' \/ T x y'
H2':T x y' -> y = y' \/ T x y
H4':T x y'' -> y' = y'' \/ T x y'
H4:T x y' -> y' = y'' \/ T x y''
H:T x y
H0:T x y'

y = y'' \/ T x y''
destruct (H2' H0), (H4 H0); try subst y'; try subst y''; try firstorder.
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y0 := snd z in let y'0 := snd z' in x0 = x' /\ (T x0 y0 -> y0 = y'0 \/ T x0 y'0) /\ (T x0 y'0 -> y0 = y'0 \/ T x0 y0):D -> D -> Prop
x:A
y, y', y'':B
H2:T x y -> y = y' \/ T x y'
H2':T x y' -> y = y' \/ T x y
H4':T x y'' -> y' = y'' \/ T x y'
H4:T x y' -> y' = y'' \/ T x y''
H:T (fst (x, y)) (snd (x, y''))

snd (x, y) = snd (x, y'') \/ T (fst (x, y)) (snd (x, y))
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y0 := snd z in let y'0 := snd z' in x0 = x' /\ (T x0 y0 -> y0 = y'0 \/ T x0 y'0) /\ (T x0 y'0 -> y0 = y'0 \/ T x0 y0):D -> D -> Prop
x:A
y, y', y'':B
H2:T x y -> y = y' \/ T x y'
H2':T x y' -> y = y' \/ T x y
H4':T x y'' -> y' = y'' \/ T x y'
H4:T x y' -> y' = y'' \/ T x y''
H:T x (snd (x, y''))

snd (x, y) = snd (x, y'') \/ T x (snd (x, y))
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y0 := snd z in let y'0 := snd z' in x0 = x' /\ (T x0 y0 -> y0 = y'0 \/ T x0 y'0) /\ (T x0 y'0 -> y0 = y'0 \/ T x0 y0):D -> D -> Prop
x:A
y, y', y'':B
H2:T x y -> y = y' \/ T x y'
H2':T x y' -> y = y' \/ T x y
H4':T x y'' -> y' = y'' \/ T x y'
H4:T x y' -> y' = y'' \/ T x y''
H:T x y''

y = y'' \/ T x y
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y0 := snd z in let y'0 := snd z' in x0 = x' /\ (T x0 y0 -> y0 = y'0 \/ T x0 y'0) /\ (T x0 y'0 -> y0 = y'0 \/ T x0 y0):D -> D -> Prop
x:A
y, y':B
H2:T x y -> y = y' \/ T x y'
H2':T x y' -> y = y' \/ T x y
H:T x y'
H4, H4':T x y' -> y' = y' \/ T x y'

y = y' \/ T x y
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y0 := snd z in let y'0 := snd z' in x0 = x' /\ (T x0 y0 -> y0 = y'0 \/ T x0 y'0) /\ (T x0 y'0 -> y0 = y'0 \/ T x0 y0):D -> D -> Prop
x:A
y, y', y'':B
H2:T x y -> y = y' \/ T x y'
H2':T x y' -> y = y' \/ T x y
H4':T x y'' -> y' = y'' \/ T x y'
H4:T x y' -> y' = y'' \/ T x y''
H:T x y''
H0:T x y'
y = y'' \/ T x y
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y0 := snd z in let y'0 := snd z' in x0 = x' /\ (T x0 y0 -> y0 = y'0 \/ T x0 y'0) /\ (T x0 y'0 -> y0 = y'0 \/ T x0 y0):D -> D -> Prop
x:A
y, y':B
H2:T x y -> y = y' \/ T x y'
H2':T x y' -> y = y' \/ T x y
H:T x y'
H4, H4':T x y' -> y' = y' \/ T x y'

y = y' \/ T x y
destruct (H2' H); firstorder.
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y0 := snd z in let y'0 := snd z' in x0 = x' /\ (T x0 y0 -> y0 = y'0 \/ T x0 y'0) /\ (T x0 y'0 -> y0 = y'0 \/ T x0 y0):D -> D -> Prop
x:A
y, y', y'':B
H2:T x y -> y = y' \/ T x y'
H2':T x y' -> y = y' \/ T x y
H4':T x y'' -> y' = y'' \/ T x y'
H4:T x y' -> y' = y'' \/ T x y''
H:T x y''
H0:T x y'

y = y'' \/ T x y
destruct (H2' H0), (H4 H0); try subst y'; try subst y''; try firstorder.
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x : A, exists y : B, T x y
D:=(A * B)%type:Type
R:=fun z z' : D => let x := fst z in let x' := fst z' in let y := snd z in let y' := snd z' in x = x' /\ (T x y -> y = y' \/ T x y') /\ (T x y' -> y = y' \/ T x y):D -> D -> Prop
Hequiv:Equivalence R

exists R' : A -> B -> Prop, subrelation R' T /\ (forall x : A, exists ! y : B, R' x y)
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x : A, exists y : B, T x y
D:=(A * B)%type:Type
R:=fun z z' : D => let x := fst z in let x' := fst z' in let y := snd z in let y' := snd z' in x = x' /\ (T x y -> y = y' \/ T x y') /\ (T x y' -> y = y' \/ T x y):D -> D -> Prop
Hequiv:Equivalence R
g:D -> D
Hg:forall x : D, R x (g x) /\ (forall x' : D, R x x' -> g x = g x')

exists R' : A -> B -> Prop, subrelation R' T /\ (forall x : A, exists ! y : B, R' x y)
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x : A, exists y : B, T x y
D:=(A * B)%type:Type
R:=fun z z' : D => let x := fst z in let x' := fst z' in let y := snd z in let y' := snd z' in x = x' /\ (T x y -> y = y' \/ T x y') /\ (T x y' -> y = y' \/ T x y):D -> D -> Prop
Hequiv:Equivalence R
g:D -> D
Hg:forall x : D, R x (g x) /\ (forall x' : D, R x x' -> g x = g x')
T':=fun (x : A) (y : B) => T x y /\ (exists y' : B, T x y' /\ g (x, y') = (x, y)):A -> B -> Prop

exists R' : A -> B -> Prop, subrelation R' T /\ (forall x : A, exists ! y : B, R' x y)
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x : A, exists y : B, T x y
D:=(A * B)%type:Type
R:=fun z z' : D => let x := fst z in let x' := fst z' in let y := snd z in let y' := snd z' in x = x' /\ (T x y -> y = y' \/ T x y') /\ (T x y' -> y = y' \/ T x y):D -> D -> Prop
Hequiv:Equivalence R
g:D -> D
Hg:forall x : D, R x (g x) /\ (forall x' : D, R x x' -> g x = g x')
T':=fun (x : A) (y : B) => T x y /\ (exists y' : B, T x y' /\ g (x, y') = (x, y)):A -> B -> Prop

subrelation T' T /\ (forall x : A, exists ! y : B, T' x y)
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x : A, exists y : B, T x y
D:=(A * B)%type:Type
R:=fun z z' : D => let x := fst z in let x' := fst z' in let y := snd z in let y' := snd z' in x = x' /\ (T x y -> y = y' \/ T x y') /\ (T x y' -> y = y' \/ T x y):D -> D -> Prop
Hequiv:Equivalence R
g:D -> D
Hg:forall x : D, R x (g x) /\ (forall x' : D, R x x' -> g x = g x')
T':=fun (x : A) (y : B) => T x y /\ (exists y' : B, T x y' /\ g (x, y') = (x, y)):A -> B -> Prop

subrelation T' T
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x : A, exists y : B, T x y
D:=(A * B)%type:Type
R:=fun z z' : D => let x := fst z in let x' := fst z' in let y := snd z in let y' := snd z' in x = x' /\ (T x y -> y = y' \/ T x y') /\ (T x y' -> y = y' \/ T x y):D -> D -> Prop
Hequiv:Equivalence R
g:D -> D
Hg:forall x : D, R x (g x) /\ (forall x' : D, R x x' -> g x = g x')
T':=fun (x : A) (y : B) => T x y /\ (exists y' : B, T x y' /\ g (x, y') = (x, y)):A -> B -> Prop
forall x : A, exists ! y : B, T' x y
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x : A, exists y : B, T x y
D:=(A * B)%type:Type
R:=fun z z' : D => let x := fst z in let x' := fst z' in let y := snd z in let y' := snd z' in x = x' /\ (T x y -> y = y' \/ T x y') /\ (T x y' -> y = y' \/ T x y):D -> D -> Prop
Hequiv:Equivalence R
g:D -> D
Hg:forall x : D, R x (g x) /\ (forall x' : D, R x x' -> g x = g x')
T':=fun (x : A) (y : B) => T x y /\ (exists y' : B, T x y' /\ g (x, y') = (x, y)):A -> B -> Prop

subrelation T' T
intros x y (H,_); easy.
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x : A, exists y : B, T x y
D:=(A * B)%type:Type
R:=fun z z' : D => let x := fst z in let x' := fst z' in let y := snd z in let y' := snd z' in x = x' /\ (T x y -> y = y' \/ T x y') /\ (T x y' -> y = y' \/ T x y):D -> D -> Prop
Hequiv:Equivalence R
g:D -> D
Hg:forall x : D, R x (g x) /\ (forall x' : D, R x x' -> g x = g x')
T':=fun (x : A) (y : B) => T x y /\ (exists y' : B, T x y' /\ g (x, y') = (x, y)):A -> B -> Prop

forall x : A, exists ! y : B, T' x y
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y : B, T x0 y
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y := snd z in let y' := snd z' in x0 = x' /\ (T x0 y -> y = y' \/ T x0 y') /\ (T x0 y' -> y = y' \/ T x0 y):D -> D -> Prop
Hequiv:Equivalence R
g:D -> D
Hg:forall x0 : D, R x0 (g x0) /\ (forall x' : D, R x0 x' -> g x0 = g x')
T':=fun (x0 : A) (y : B) => T x0 y /\ (exists y' : B, T x0 y' /\ g (x0, y') = (x0, y)):A -> B -> Prop
x:A

exists ! y : B, T' x y
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y0 := snd z in let y' := snd z' in x0 = x' /\ (T x0 y0 -> y0 = y' \/ T x0 y') /\ (T x0 y' -> y0 = y' \/ T x0 y0):D -> D -> Prop
Hequiv:Equivalence R
g:D -> D
Hg:forall x0 : D, R x0 (g x0) /\ (forall x' : D, R x0 x' -> g x0 = g x')
T':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y' : B, T x0 y' /\ g (x0, y') = (x0, y0)):A -> B -> Prop
x:A
y:B
Hy:T x y

exists ! y0 : B, T' x y0
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y0 := snd z in let y' := snd z' in x0 = x' /\ (T x0 y0 -> y0 = y' \/ T x0 y') /\ (T x0 y' -> y0 = y' \/ T x0 y0):D -> D -> Prop
Hequiv:Equivalence R
g:D -> D
Hg:forall x0 : D, R x0 (g x0) /\ (forall x' : D, R x0 x' -> g x0 = g x')
T':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y' : B, T x0 y' /\ g (x0, y') = (x0, y0)):A -> B -> Prop
x:A
y:B
Hy:T x y

unique (fun y0 : B => T' x y0) (snd (g (x, y)))
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y0 := snd z in let y' := snd z' in x0 = x' /\ (T x0 y0 -> y0 = y' \/ T x0 y') /\ (T x0 y' -> y0 = y' \/ T x0 y0):D -> D -> Prop
Hequiv:Equivalence R
g:D -> D
T':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y' : B, T x0 y' /\ g (x0, y') = (x0, y0)):A -> B -> Prop
x:A
y:B
Hy:T x y
Heq1:fst (x, y) = fst (g (x, y))
H':T (fst (x, y)) (snd (x, y)) -> snd (x, y) = snd (g (x, y)) \/ T (fst (x, y)) (snd (g (x, y)))
H'':T (fst (x, y)) (snd (g (x, y))) -> snd (x, y) = snd (g (x, y)) \/ T (fst (x, y)) (snd (x, y))
Hgxyuniq:forall x' : D, R (x, y) x' -> g (x, y) = g x'

unique (fun y0 : B => T' x y0) (snd (g (x, y)))
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y0 := snd z in let y' := snd z' in x0 = x' /\ (T x0 y0 -> y0 = y' \/ T x0 y') /\ (T x0 y' -> y0 = y' \/ T x0 y0):D -> D -> Prop
Hequiv:Equivalence R
g:D -> D
T':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y' : B, T x0 y' /\ g (x0, y') = (x0, y0)):A -> B -> Prop
x:A
y:B
Hy:T x y
Heq1:fst (x, y) = fst (g (x, y))
H'':T (fst (x, y)) (snd (g (x, y))) -> snd (x, y) = snd (g (x, y)) \/ T (fst (x, y)) (snd (x, y))
Hgxyuniq:forall x' : D, R (x, y) x' -> g (x, y) = g x'
Heq2:snd (x, y) = snd (g (x, y))

unique (fun y0 : B => T' x y0) (snd (g (x, y)))
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y0 := snd z in let y' := snd z' in x0 = x' /\ (T x0 y0 -> y0 = y' \/ T x0 y') /\ (T x0 y' -> y0 = y' \/ T x0 y0):D -> D -> Prop
Hequiv:Equivalence R
g:D -> D
T':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y' : B, T x0 y' /\ g (x0, y') = (x0, y0)):A -> B -> Prop
x:A
y:B
Hy:T x y
Heq1:fst (x, y) = fst (g (x, y))
H'':T (fst (x, y)) (snd (g (x, y))) -> snd (x, y) = snd (g (x, y)) \/ T (fst (x, y)) (snd (x, y))
Hgxyuniq:forall x' : D, R (x, y) x' -> g (x, y) = g x'
Hgy:T (fst (x, y)) (snd (g (x, y)))
unique (fun y0 : B => T' x y0) (snd (g (x, y)))
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y0 := snd z in let y' := snd z' in x0 = x' /\ (T x0 y0 -> y0 = y' \/ T x0 y') /\ (T x0 y' -> y0 = y' \/ T x0 y0):D -> D -> Prop
Hequiv:Equivalence R
g:D -> D
T':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y' : B, T x0 y' /\ g (x0, y') = (x0, y0)):A -> B -> Prop
x:A
y:B
Hy:T x y
Heq1:fst (x, y) = fst (g (x, y))
H'':T (fst (x, y)) (snd (g (x, y))) -> snd (x, y) = snd (g (x, y)) \/ T (fst (x, y)) (snd (x, y))
Hgxyuniq:forall x' : D, R (x, y) x' -> g (x, y) = g x'
Heq2:snd (x, y) = snd (g (x, y))

unique (fun y0 : B => T' x y0) (snd (g (x, y)))
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y0 := snd z in let y' := snd z' in x0 = x' /\ (T x0 y0 -> y0 = y' \/ T x0 y') /\ (T x0 y' -> y0 = y' \/ T x0 y0):D -> D -> Prop
Hequiv:Equivalence R
g:D -> D
T':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y' : B, T x0 y' /\ g (x0, y') = (x0, y0)):A -> B -> Prop
x:A
y:B
Hy:T x y
Heq1:fst (x, y) = fst (g (x, y))
H'':T (fst (x, y)) (snd (g (x, y))) -> snd (x, y) = snd (g (x, y)) \/ T (fst (x, y)) (snd (x, y))
Hgxyuniq:forall x' : D, R (x, y) x' -> g (x, y) = g x'
Heq2:snd (x, y) = snd (g (x, y))

T' x (snd (g (x, y)))
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y0 := snd z in let y' := snd z' in x0 = x' /\ (T x0 y0 -> y0 = y' \/ T x0 y') /\ (T x0 y' -> y0 = y' \/ T x0 y0):D -> D -> Prop
Hequiv:Equivalence R
g:D -> D
T':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y' : B, T x0 y' /\ g (x0, y') = (x0, y0)):A -> B -> Prop
x:A
y:B
Hy:T x y
Heq1:fst (x, y) = fst (g (x, y))
H'':T (fst (x, y)) (snd (g (x, y))) -> snd (x, y) = snd (g (x, y)) \/ T (fst (x, y)) (snd (x, y))
Hgxyuniq:forall x' : D, R (x, y) x' -> g (x, y) = g x'
Heq2:snd (x, y) = snd (g (x, y))
forall x' : B, T' x x' -> snd (g (x, y)) = x'
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y0 := snd z in let y' := snd z' in x0 = x' /\ (T x0 y0 -> y0 = y' \/ T x0 y') /\ (T x0 y' -> y0 = y' \/ T x0 y0):D -> D -> Prop
Hequiv:Equivalence R
g:D -> D
T':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y' : B, T x0 y' /\ g (x0, y') = (x0, y0)):A -> B -> Prop
x:A
y:B
Hy:T x y
Heq1:fst (x, y) = fst (g (x, y))
H'':T (fst (x, y)) (snd (g (x, y))) -> snd (x, y) = snd (g (x, y)) \/ T (fst (x, y)) (snd (x, y))
Hgxyuniq:forall x' : D, R (x, y) x' -> g (x, y) = g x'
Heq2:snd (x, y) = snd (g (x, y))

T x (snd (g (x, y)))
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y0 := snd z in let y' := snd z' in x0 = x' /\ (T x0 y0 -> y0 = y' \/ T x0 y') /\ (T x0 y' -> y0 = y' \/ T x0 y0):D -> D -> Prop
Hequiv:Equivalence R
g:D -> D
T':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y' : B, T x0 y' /\ g (x0, y') = (x0, y0)):A -> B -> Prop
x:A
y:B
Hy:T x y
Heq1:fst (x, y) = fst (g (x, y))
H'':T (fst (x, y)) (snd (g (x, y))) -> snd (x, y) = snd (g (x, y)) \/ T (fst (x, y)) (snd (x, y))
Hgxyuniq:forall x' : D, R (x, y) x' -> g (x, y) = g x'
Heq2:snd (x, y) = snd (g (x, y))
exists y' : B, T x y' /\ g (x, y') = (x, snd (g (x, y)))
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y0 := snd z in let y' := snd z' in x0 = x' /\ (T x0 y0 -> y0 = y' \/ T x0 y') /\ (T x0 y' -> y0 = y' \/ T x0 y0):D -> D -> Prop
Hequiv:Equivalence R
g:D -> D
T':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y' : B, T x0 y' /\ g (x0, y') = (x0, y0)):A -> B -> Prop
x:A
y:B
Hy:T x y
Heq1:fst (x, y) = fst (g (x, y))
H'':T (fst (x, y)) (snd (g (x, y))) -> snd (x, y) = snd (g (x, y)) \/ T (fst (x, y)) (snd (x, y))
Hgxyuniq:forall x' : D, R (x, y) x' -> g (x, y) = g x'
Heq2:snd (x, y) = snd (g (x, y))
forall x' : B, T' x x' -> snd (g (x, y)) = x'
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y0 := snd z in let y' := snd z' in x0 = x' /\ (T x0 y0 -> y0 = y' \/ T x0 y') /\ (T x0 y' -> y0 = y' \/ T x0 y0):D -> D -> Prop
Hequiv:Equivalence R
g:D -> D
T':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y' : B, T x0 y' /\ g (x0, y') = (x0, y0)):A -> B -> Prop
x:A
y:B
Hy:T x y
Heq1:fst (x, y) = fst (g (x, y))
H'':T (fst (x, y)) (snd (g (x, y))) -> snd (x, y) = snd (g (x, y)) \/ T (fst (x, y)) (snd (x, y))
Hgxyuniq:forall x' : D, R (x, y) x' -> g (x, y) = g x'
Heq2:snd (x, y) = snd (g (x, y))

T x (snd (g (x, y)))
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y0 := snd z in let y' := snd z' in x0 = x' /\ (T x0 y0 -> y0 = y' \/ T x0 y') /\ (T x0 y' -> y0 = y' \/ T x0 y0):D -> D -> Prop
Hequiv:Equivalence R
g:D -> D
T':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y' : B, T x0 y' /\ g (x0, y') = (x0, y0)):A -> B -> Prop
x:A
y:B
Hy:T x y
Heq1:fst (x, y) = fst (g (x, y))
H'':T (fst (x, y)) (snd (g (x, y))) -> snd (x, y) = snd (g (x, y)) \/ T (fst (x, y)) (snd (x, y))
Hgxyuniq:forall x' : D, R (x, y) x' -> g (x, y) = g x'
Heq2:snd (x, y) = snd (g (x, y))

T x (snd (x, y))
assumption.
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y0 := snd z in let y' := snd z' in x0 = x' /\ (T x0 y0 -> y0 = y' \/ T x0 y') /\ (T x0 y' -> y0 = y' \/ T x0 y0):D -> D -> Prop
Hequiv:Equivalence R
g:D -> D
T':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y' : B, T x0 y' /\ g (x0, y') = (x0, y0)):A -> B -> Prop
x:A
y:B
Hy:T x y
Heq1:fst (x, y) = fst (g (x, y))
H'':T (fst (x, y)) (snd (g (x, y))) -> snd (x, y) = snd (g (x, y)) \/ T (fst (x, y)) (snd (x, y))
Hgxyuniq:forall x' : D, R (x, y) x' -> g (x, y) = g x'
Heq2:snd (x, y) = snd (g (x, y))

exists y' : B, T x y' /\ g (x, y') = (x, snd (g (x, y)))
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y0 := snd z in let y' := snd z' in x0 = x' /\ (T x0 y0 -> y0 = y' \/ T x0 y') /\ (T x0 y' -> y0 = y' \/ T x0 y0):D -> D -> Prop
Hequiv:Equivalence R
g:D -> D
T':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y' : B, T x0 y' /\ g (x0, y') = (x0, y0)):A -> B -> Prop
x:A
y:B
Hy:T x y
Heq1:fst (x, y) = fst (g (x, y))
H'':T (fst (x, y)) (snd (g (x, y))) -> snd (x, y) = snd (g (x, y)) \/ T (fst (x, y)) (snd (x, y))
Hgxyuniq:forall x' : D, R (x, y) x' -> g (x, y) = g x'
Heq2:snd (x, y) = snd (g (x, y))

T x y /\ g (x, y) = (x, snd (g (x, y)))
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x'0 := fst z' in let y0 := snd z in let y'0 := snd z' in x0 = x'0 /\ (T x0 y0 -> y0 = y'0 \/ T x0 y'0) /\ (T x0 y'0 -> y0 = y'0 \/ T x0 y0):D -> D -> Prop
Hequiv:Equivalence R
g:D -> D
T':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y'0 : B, T x0 y'0 /\ g (x0, y'0) = (x0, y0)):A -> B -> Prop
x:A
y:B
Hy:T x y
x':A
y':B
Heq1:fst (x, y) = fst (x', y')
H'':T (fst (x, y)) (snd (x', y')) -> snd (x, y) = snd (x', y') \/ T (fst (x, y)) (snd (x, y))
Hgxyuniq:forall x'0 : D, R (x, y) x'0 -> (x', y') = g x'0
Heq2:snd (x, y) = snd (x', y')

T x y /\ (x', y') = (x, snd (x', y'))
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x'0 := fst z' in let y0 := snd z in let y'0 := snd z' in x0 = x'0 /\ (T x0 y0 -> y0 = y'0 \/ T x0 y'0) /\ (T x0 y'0 -> y0 = y'0 \/ T x0 y0):D -> D -> Prop
Hequiv:Equivalence R
g:D -> D
T':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y'0 : B, T x0 y'0 /\ g (x0, y'0) = (x0, y0)):A -> B -> Prop
x:A
y:B
Hy:T x y
x':A
y':B
Heq1:x = x'
H'':T (fst (x, y)) (snd (x', y')) -> snd (x, y) = snd (x', y') \/ T (fst (x, y)) (snd (x, y))
Hgxyuniq:forall x'0 : D, R (x, y) x'0 -> (x', y') = g x'0
Heq2:y = y'

T x y /\ (x', y') = (x, snd (x', y'))
subst; easy.
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y0 := snd z in let y' := snd z' in x0 = x' /\ (T x0 y0 -> y0 = y' \/ T x0 y') /\ (T x0 y' -> y0 = y' \/ T x0 y0):D -> D -> Prop
Hequiv:Equivalence R
g:D -> D
T':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y' : B, T x0 y' /\ g (x0, y') = (x0, y0)):A -> B -> Prop
x:A
y:B
Hy:T x y
Heq1:fst (x, y) = fst (g (x, y))
H'':T (fst (x, y)) (snd (g (x, y))) -> snd (x, y) = snd (g (x, y)) \/ T (fst (x, y)) (snd (x, y))
Hgxyuniq:forall x' : D, R (x, y) x' -> g (x, y) = g x'
Heq2:snd (x, y) = snd (g (x, y))

forall x' : B, T' x x' -> snd (g (x, y)) = x'
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y0 := snd z in let y'0 := snd z' in x0 = x' /\ (T x0 y0 -> y0 = y'0 \/ T x0 y'0) /\ (T x0 y'0 -> y0 = y'0 \/ T x0 y0):D -> D -> Prop
Hequiv:Equivalence R
g:D -> D
T':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y'0 : B, T x0 y'0 /\ g (x0, y'0) = (x0, y0)):A -> B -> Prop
x:A
y:B
Hy:T x y
Heq1:fst (x, y) = fst (g (x, y))
H'':T (fst (x, y)) (snd (g (x, y))) -> snd (x, y) = snd (g (x, y)) \/ T (fst (x, y)) (snd (x, y))
Hgxyuniq:forall x' : D, R (x, y) x' -> g (x, y) = g x'
Heq2:snd (x, y) = snd (g (x, y))
y':B
Hy':T x y'
y'':B
Hy'':T x y''
Heq:g (x, y'') = (x, y')

snd (g (x, y)) = y'
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y0 := snd z in let y'0 := snd z' in x0 = x' /\ (T x0 y0 -> y0 = y'0 \/ T x0 y'0) /\ (T x0 y'0 -> y0 = y'0 \/ T x0 y0):D -> D -> Prop
Hequiv:Equivalence R
g:D -> D
T':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y'0 : B, T x0 y'0 /\ g (x0, y'0) = (x0, y0)):A -> B -> Prop
x:A
y:B
Hy:T x y
Heq1:fst (x, y) = fst (g (x, y))
H'':T (fst (x, y)) (snd (g (x, y))) -> snd (x, y) = snd (g (x, y)) \/ T (fst (x, y)) (snd (x, y))
Hgxyuniq:forall x' : D, R (x, y) x' -> g (x, y) = g x'
Heq2:snd (x, y) = snd (g (x, y))
y':B
Hy':T x y'
y'':B
Hy'':T x y''
Heq:g (x, y'') = (x, y')

snd (x, y') = y'
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y0 := snd z in let y'0 := snd z' in x0 = x' /\ (T x0 y0 -> y0 = y'0 \/ T x0 y'0) /\ (T x0 y'0 -> y0 = y'0 \/ T x0 y0):D -> D -> Prop
Hequiv:Equivalence R
g:D -> D
T':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y'0 : B, T x0 y'0 /\ g (x0, y'0) = (x0, y0)):A -> B -> Prop
x:A
y:B
Hy:T x y
Heq1:fst (x, y) = fst (g (x, y))
H'':T (fst (x, y)) (snd (g (x, y))) -> snd (x, y) = snd (g (x, y)) \/ T (fst (x, y)) (snd (x, y))
Hgxyuniq:forall x' : D, R (x, y) x' -> g (x, y) = g x'
Heq2:snd (x, y) = snd (g (x, y))
y':B
Hy':T x y'
y'':B
Hy'':T x y''
Heq:g (x, y'') = (x, y')
R (x, y) (x, y'')
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y0 := snd z in let y'0 := snd z' in x0 = x' /\ (T x0 y0 -> y0 = y'0 \/ T x0 y'0) /\ (T x0 y'0 -> y0 = y'0 \/ T x0 y0):D -> D -> Prop
Hequiv:Equivalence R
g:D -> D
T':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y'0 : B, T x0 y'0 /\ g (x0, y'0) = (x0, y0)):A -> B -> Prop
x:A
y:B
Hy:T x y
Heq1:fst (x, y) = fst (g (x, y))
H'':T (fst (x, y)) (snd (g (x, y))) -> snd (x, y) = snd (g (x, y)) \/ T (fst (x, y)) (snd (x, y))
Hgxyuniq:forall x' : D, R (x, y) x' -> g (x, y) = g x'
Heq2:snd (x, y) = snd (g (x, y))
y':B
Hy':T x y'
y'':B
Hy'':T x y''
Heq:g (x, y'') = (x, y')

R (x, y) (x, y'')
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y0 := snd z in let y'0 := snd z' in x0 = x' /\ (T x0 y0 -> y0 = y'0 \/ T x0 y'0) /\ (T x0 y'0 -> y0 = y'0 \/ T x0 y0):D -> D -> Prop
Hequiv:Equivalence R
g:D -> D
T':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y'0 : B, T x0 y'0 /\ g (x0, y'0) = (x0, y0)):A -> B -> Prop
x:A
y:B
Hy:T x y
Heq1:fst (x, y) = fst (g (x, y))
H'':T (fst (x, y)) (snd (g (x, y))) -> snd (x, y) = snd (g (x, y)) \/ T (fst (x, y)) (snd (x, y))
Hgxyuniq:forall x' : D, R (x, y) x' -> g (x, y) = g x'
Heq2:snd (x, y) = snd (g (x, y))
y':B
Hy':T x y'
y'':B
Hy'':T x y''
Heq:g (x, y'') = (x, y')

fst (x, y) = fst (x, y'')
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y0 := snd z in let y'0 := snd z' in x0 = x' /\ (T x0 y0 -> y0 = y'0 \/ T x0 y'0) /\ (T x0 y'0 -> y0 = y'0 \/ T x0 y0):D -> D -> Prop
Hequiv:Equivalence R
g:D -> D
T':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y'0 : B, T x0 y'0 /\ g (x0, y'0) = (x0, y0)):A -> B -> Prop
x:A
y:B
Hy:T x y
Heq1:fst (x, y) = fst (g (x, y))
H'':T (fst (x, y)) (snd (g (x, y))) -> snd (x, y) = snd (g (x, y)) \/ T (fst (x, y)) (snd (x, y))
Hgxyuniq:forall x' : D, R (x, y) x' -> g (x, y) = g x'
Heq2:snd (x, y) = snd (g (x, y))
y':B
Hy':T x y'
y'':B
Hy'':T x y''
Heq:g (x, y'') = (x, y')
(T (fst (x, y)) (snd (x, y)) -> snd (x, y) = snd (x, y'') \/ T (fst (x, y)) (snd (x, y''))) /\ (T (fst (x, y)) (snd (x, y'')) -> snd (x, y) = snd (x, y'') \/ T (fst (x, y)) (snd (x, y)))
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y0 := snd z in let y'0 := snd z' in x0 = x' /\ (T x0 y0 -> y0 = y'0 \/ T x0 y'0) /\ (T x0 y'0 -> y0 = y'0 \/ T x0 y0):D -> D -> Prop
Hequiv:Equivalence R
g:D -> D
T':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y'0 : B, T x0 y'0 /\ g (x0, y'0) = (x0, y0)):A -> B -> Prop
x:A
y:B
Hy:T x y
Heq1:fst (x, y) = fst (g (x, y))
H'':T (fst (x, y)) (snd (g (x, y))) -> snd (x, y) = snd (g (x, y)) \/ T (fst (x, y)) (snd (x, y))
Hgxyuniq:forall x' : D, R (x, y) x' -> g (x, y) = g x'
Heq2:snd (x, y) = snd (g (x, y))
y':B
Hy':T x y'
y'':B
Hy'':T x y''
Heq:g (x, y'') = (x, y')

(T (fst (x, y)) (snd (x, y)) -> snd (x, y) = snd (x, y'') \/ T (fst (x, y)) (snd (x, y''))) /\ (T (fst (x, y)) (snd (x, y'')) -> snd (x, y) = snd (x, y'') \/ T (fst (x, y)) (snd (x, y)))
split; right; easy.
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y0 := snd z in let y' := snd z' in x0 = x' /\ (T x0 y0 -> y0 = y' \/ T x0 y') /\ (T x0 y' -> y0 = y' \/ T x0 y0):D -> D -> Prop
Hequiv:Equivalence R
g:D -> D
T':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y' : B, T x0 y' /\ g (x0, y') = (x0, y0)):A -> B -> Prop
x:A
y:B
Hy:T x y
Heq1:fst (x, y) = fst (g (x, y))
H'':T (fst (x, y)) (snd (g (x, y))) -> snd (x, y) = snd (g (x, y)) \/ T (fst (x, y)) (snd (x, y))
Hgxyuniq:forall x' : D, R (x, y) x' -> g (x, y) = g x'
Hgy:T (fst (x, y)) (snd (g (x, y)))

unique (fun y0 : B => T' x y0) (snd (g (x, y)))
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y0 := snd z in let y' := snd z' in x0 = x' /\ (T x0 y0 -> y0 = y' \/ T x0 y') /\ (T x0 y' -> y0 = y' \/ T x0 y0):D -> D -> Prop
Hequiv:Equivalence R
g:D -> D
T':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y' : B, T x0 y' /\ g (x0, y') = (x0, y0)):A -> B -> Prop
x:A
y:B
Hy:T x y
Heq1:fst (x, y) = fst (g (x, y))
H'':T (fst (x, y)) (snd (g (x, y))) -> snd (x, y) = snd (g (x, y)) \/ T (fst (x, y)) (snd (x, y))
Hgxyuniq:forall x' : D, R (x, y) x' -> g (x, y) = g x'
Hgy:T (fst (x, y)) (snd (g (x, y)))

T' x (snd (g (x, y)))
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y0 := snd z in let y' := snd z' in x0 = x' /\ (T x0 y0 -> y0 = y' \/ T x0 y') /\ (T x0 y' -> y0 = y' \/ T x0 y0):D -> D -> Prop
Hequiv:Equivalence R
g:D -> D
T':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y' : B, T x0 y' /\ g (x0, y') = (x0, y0)):A -> B -> Prop
x:A
y:B
Hy:T x y
Heq1:fst (x, y) = fst (g (x, y))
H'':T (fst (x, y)) (snd (g (x, y))) -> snd (x, y) = snd (g (x, y)) \/ T (fst (x, y)) (snd (x, y))
Hgxyuniq:forall x' : D, R (x, y) x' -> g (x, y) = g x'
Hgy:T (fst (x, y)) (snd (g (x, y)))
forall x' : B, T' x x' -> snd (g (x, y)) = x'
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y0 := snd z in let y' := snd z' in x0 = x' /\ (T x0 y0 -> y0 = y' \/ T x0 y') /\ (T x0 y' -> y0 = y' \/ T x0 y0):D -> D -> Prop
Hequiv:Equivalence R
g:D -> D
T':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y' : B, T x0 y' /\ g (x0, y') = (x0, y0)):A -> B -> Prop
x:A
y:B
Hy:T x y
Heq1:fst (x, y) = fst (g (x, y))
H'':T (fst (x, y)) (snd (g (x, y))) -> snd (x, y) = snd (g (x, y)) \/ T (fst (x, y)) (snd (x, y))
Hgxyuniq:forall x' : D, R (x, y) x' -> g (x, y) = g x'
Hgy:T (fst (x, y)) (snd (g (x, y)))

T x (snd (g (x, y)))
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y0 := snd z in let y' := snd z' in x0 = x' /\ (T x0 y0 -> y0 = y' \/ T x0 y') /\ (T x0 y' -> y0 = y' \/ T x0 y0):D -> D -> Prop
Hequiv:Equivalence R
g:D -> D
T':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y' : B, T x0 y' /\ g (x0, y') = (x0, y0)):A -> B -> Prop
x:A
y:B
Hy:T x y
Heq1:fst (x, y) = fst (g (x, y))
H'':T (fst (x, y)) (snd (g (x, y))) -> snd (x, y) = snd (g (x, y)) \/ T (fst (x, y)) (snd (x, y))
Hgxyuniq:forall x' : D, R (x, y) x' -> g (x, y) = g x'
Hgy:T (fst (x, y)) (snd (g (x, y)))
exists y' : B, T x y' /\ g (x, y') = (x, snd (g (x, y)))
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y0 := snd z in let y' := snd z' in x0 = x' /\ (T x0 y0 -> y0 = y' \/ T x0 y') /\ (T x0 y' -> y0 = y' \/ T x0 y0):D -> D -> Prop
Hequiv:Equivalence R
g:D -> D
T':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y' : B, T x0 y' /\ g (x0, y') = (x0, y0)):A -> B -> Prop
x:A
y:B
Hy:T x y
Heq1:fst (x, y) = fst (g (x, y))
H'':T (fst (x, y)) (snd (g (x, y))) -> snd (x, y) = snd (g (x, y)) \/ T (fst (x, y)) (snd (x, y))
Hgxyuniq:forall x' : D, R (x, y) x' -> g (x, y) = g x'
Hgy:T (fst (x, y)) (snd (g (x, y)))
forall x' : B, T' x x' -> snd (g (x, y)) = x'
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y0 := snd z in let y' := snd z' in x0 = x' /\ (T x0 y0 -> y0 = y' \/ T x0 y') /\ (T x0 y' -> y0 = y' \/ T x0 y0):D -> D -> Prop
Hequiv:Equivalence R
g:D -> D
T':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y' : B, T x0 y' /\ g (x0, y') = (x0, y0)):A -> B -> Prop
x:A
y:B
Hy:T x y
Heq1:fst (x, y) = fst (g (x, y))
H'':T (fst (x, y)) (snd (g (x, y))) -> snd (x, y) = snd (g (x, y)) \/ T (fst (x, y)) (snd (x, y))
Hgxyuniq:forall x' : D, R (x, y) x' -> g (x, y) = g x'
Hgy:T (fst (x, y)) (snd (g (x, y)))

T x (snd (g (x, y)))
assumption.
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y0 := snd z in let y' := snd z' in x0 = x' /\ (T x0 y0 -> y0 = y' \/ T x0 y') /\ (T x0 y' -> y0 = y' \/ T x0 y0):D -> D -> Prop
Hequiv:Equivalence R
g:D -> D
T':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y' : B, T x0 y' /\ g (x0, y') = (x0, y0)):A -> B -> Prop
x:A
y:B
Hy:T x y
Heq1:fst (x, y) = fst (g (x, y))
H'':T (fst (x, y)) (snd (g (x, y))) -> snd (x, y) = snd (g (x, y)) \/ T (fst (x, y)) (snd (x, y))
Hgxyuniq:forall x' : D, R (x, y) x' -> g (x, y) = g x'
Hgy:T (fst (x, y)) (snd (g (x, y)))

exists y' : B, T x y' /\ g (x, y') = (x, snd (g (x, y)))
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y0 := snd z in let y' := snd z' in x0 = x' /\ (T x0 y0 -> y0 = y' \/ T x0 y') /\ (T x0 y' -> y0 = y' \/ T x0 y0):D -> D -> Prop
Hequiv:Equivalence R
g:D -> D
T':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y' : B, T x0 y' /\ g (x0, y') = (x0, y0)):A -> B -> Prop
x:A
y:B
Hy:T x y
Heq1:fst (x, y) = fst (g (x, y))
H'':T (fst (x, y)) (snd (g (x, y))) -> snd (x, y) = snd (g (x, y)) \/ T (fst (x, y)) (snd (x, y))
Hgxyuniq:forall x' : D, R (x, y) x' -> g (x, y) = g x'
Hgy:T (fst (x, y)) (snd (g (x, y)))

T x y /\ g (x, y) = (x, snd (g (x, y)))
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x'0 := fst z' in let y0 := snd z in let y'0 := snd z' in x0 = x'0 /\ (T x0 y0 -> y0 = y'0 \/ T x0 y'0) /\ (T x0 y'0 -> y0 = y'0 \/ T x0 y0):D -> D -> Prop
Hequiv:Equivalence R
g:D -> D
T':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y'0 : B, T x0 y'0 /\ g (x0, y'0) = (x0, y0)):A -> B -> Prop
x:A
y:B
Hy:T x y
x':A
y':B
Heq1:fst (x, y) = fst (x', y')
H'':T (fst (x, y)) (snd (x', y')) -> snd (x, y) = snd (x', y') \/ T (fst (x, y)) (snd (x, y))
Hgxyuniq:forall x'0 : D, R (x, y) x'0 -> (x', y') = g x'0
Hgy:T (fst (x, y)) (snd (x', y'))

T x y /\ (x', y') = (x, snd (x', y'))
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x'0 := fst z' in let y0 := snd z in let y'0 := snd z' in x0 = x'0 /\ (T x0 y0 -> y0 = y'0 \/ T x0 y'0) /\ (T x0 y'0 -> y0 = y'0 \/ T x0 y0):D -> D -> Prop
Hequiv:Equivalence R
g:D -> D
T':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y'0 : B, T x0 y'0 /\ g (x0, y'0) = (x0, y0)):A -> B -> Prop
x:A
y:B
Hy:T x y
x':A
y':B
Heq1:x = x'
H'':T (fst (x, y)) (snd (x', y')) -> snd (x, y) = snd (x', y') \/ T (fst (x, y)) (snd (x, y))
Hgxyuniq:forall x'0 : D, R (x, y) x'0 -> (x', y') = g x'0
Hgy:T (fst (x, y)) (snd (x', y'))

T x y /\ (x', y') = (x, snd (x', y'))
subst x'; easy.
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y0 := snd z in let y' := snd z' in x0 = x' /\ (T x0 y0 -> y0 = y' \/ T x0 y') /\ (T x0 y' -> y0 = y' \/ T x0 y0):D -> D -> Prop
Hequiv:Equivalence R
g:D -> D
T':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y' : B, T x0 y' /\ g (x0, y') = (x0, y0)):A -> B -> Prop
x:A
y:B
Hy:T x y
Heq1:fst (x, y) = fst (g (x, y))
H'':T (fst (x, y)) (snd (g (x, y))) -> snd (x, y) = snd (g (x, y)) \/ T (fst (x, y)) (snd (x, y))
Hgxyuniq:forall x' : D, R (x, y) x' -> g (x, y) = g x'
Hgy:T (fst (x, y)) (snd (g (x, y)))

forall x' : B, T' x x' -> snd (g (x, y)) = x'
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y0 := snd z in let y'0 := snd z' in x0 = x' /\ (T x0 y0 -> y0 = y'0 \/ T x0 y'0) /\ (T x0 y'0 -> y0 = y'0 \/ T x0 y0):D -> D -> Prop
Hequiv:Equivalence R
g:D -> D
T':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y'0 : B, T x0 y'0 /\ g (x0, y'0) = (x0, y0)):A -> B -> Prop
x:A
y:B
Hy:T x y
Heq1:fst (x, y) = fst (g (x, y))
H'':T (fst (x, y)) (snd (g (x, y))) -> snd (x, y) = snd (g (x, y)) \/ T (fst (x, y)) (snd (x, y))
Hgxyuniq:forall x' : D, R (x, y) x' -> g (x, y) = g x'
Hgy:T (fst (x, y)) (snd (g (x, y)))
y':B
Hy':T x y'
y'':B
Hy'':T x y''
Heq:g (x, y'') = (x, y')

snd (g (x, y)) = y'
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y0 := snd z in let y'0 := snd z' in x0 = x' /\ (T x0 y0 -> y0 = y'0 \/ T x0 y'0) /\ (T x0 y'0 -> y0 = y'0 \/ T x0 y0):D -> D -> Prop
Hequiv:Equivalence R
g:D -> D
T':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y'0 : B, T x0 y'0 /\ g (x0, y'0) = (x0, y0)):A -> B -> Prop
x:A
y:B
Hy:T x y
Heq1:fst (x, y) = fst (g (x, y))
H'':T (fst (x, y)) (snd (g (x, y))) -> snd (x, y) = snd (g (x, y)) \/ T (fst (x, y)) (snd (x, y))
Hgxyuniq:forall x' : D, R (x, y) x' -> g (x, y) = g x'
Hgy:T (fst (x, y)) (snd (g (x, y)))
y':B
Hy':T x y'
y'':B
Hy'':T x y''
Heq:g (x, y'') = (x, y')

snd (x, y') = y'
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y0 := snd z in let y'0 := snd z' in x0 = x' /\ (T x0 y0 -> y0 = y'0 \/ T x0 y'0) /\ (T x0 y'0 -> y0 = y'0 \/ T x0 y0):D -> D -> Prop
Hequiv:Equivalence R
g:D -> D
T':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y'0 : B, T x0 y'0 /\ g (x0, y'0) = (x0, y0)):A -> B -> Prop
x:A
y:B
Hy:T x y
Heq1:fst (x, y) = fst (g (x, y))
H'':T (fst (x, y)) (snd (g (x, y))) -> snd (x, y) = snd (g (x, y)) \/ T (fst (x, y)) (snd (x, y))
Hgxyuniq:forall x' : D, R (x, y) x' -> g (x, y) = g x'
Hgy:T (fst (x, y)) (snd (g (x, y)))
y':B
Hy':T x y'
y'':B
Hy'':T x y''
Heq:g (x, y'') = (x, y')
R (x, y) (x, y'')
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y0 := snd z in let y'0 := snd z' in x0 = x' /\ (T x0 y0 -> y0 = y'0 \/ T x0 y'0) /\ (T x0 y'0 -> y0 = y'0 \/ T x0 y0):D -> D -> Prop
Hequiv:Equivalence R
g:D -> D
T':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y'0 : B, T x0 y'0 /\ g (x0, y'0) = (x0, y0)):A -> B -> Prop
x:A
y:B
Hy:T x y
Heq1:fst (x, y) = fst (g (x, y))
H'':T (fst (x, y)) (snd (g (x, y))) -> snd (x, y) = snd (g (x, y)) \/ T (fst (x, y)) (snd (x, y))
Hgxyuniq:forall x' : D, R (x, y) x' -> g (x, y) = g x'
Hgy:T (fst (x, y)) (snd (g (x, y)))
y':B
Hy':T x y'
y'':B
Hy'':T x y''
Heq:g (x, y'') = (x, y')

R (x, y) (x, y'')
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y0 := snd z in let y'0 := snd z' in x0 = x' /\ (T x0 y0 -> y0 = y'0 \/ T x0 y'0) /\ (T x0 y'0 -> y0 = y'0 \/ T x0 y0):D -> D -> Prop
Hequiv:Equivalence R
g:D -> D
T':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y'0 : B, T x0 y'0 /\ g (x0, y'0) = (x0, y0)):A -> B -> Prop
x:A
y:B
Hy:T x y
Heq1:fst (x, y) = fst (g (x, y))
H'':T (fst (x, y)) (snd (g (x, y))) -> snd (x, y) = snd (g (x, y)) \/ T (fst (x, y)) (snd (x, y))
Hgxyuniq:forall x' : D, R (x, y) x' -> g (x, y) = g x'
Hgy:T (fst (x, y)) (snd (g (x, y)))
y':B
Hy':T x y'
y'':B
Hy'':T x y''
Heq:g (x, y'') = (x, y')

fst (x, y) = fst (x, y'')
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y0 := snd z in let y'0 := snd z' in x0 = x' /\ (T x0 y0 -> y0 = y'0 \/ T x0 y'0) /\ (T x0 y'0 -> y0 = y'0 \/ T x0 y0):D -> D -> Prop
Hequiv:Equivalence R
g:D -> D
T':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y'0 : B, T x0 y'0 /\ g (x0, y'0) = (x0, y0)):A -> B -> Prop
x:A
y:B
Hy:T x y
Heq1:fst (x, y) = fst (g (x, y))
H'':T (fst (x, y)) (snd (g (x, y))) -> snd (x, y) = snd (g (x, y)) \/ T (fst (x, y)) (snd (x, y))
Hgxyuniq:forall x' : D, R (x, y) x' -> g (x, y) = g x'
Hgy:T (fst (x, y)) (snd (g (x, y)))
y':B
Hy':T x y'
y'':B
Hy'':T x y''
Heq:g (x, y'') = (x, y')
(T (fst (x, y)) (snd (x, y)) -> snd (x, y) = snd (x, y'') \/ T (fst (x, y)) (snd (x, y''))) /\ (T (fst (x, y)) (snd (x, y'')) -> snd (x, y) = snd (x, y'') \/ T (fst (x, y)) (snd (x, y)))
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
T:A -> B -> Prop
Hexists:forall x0 : A, exists y0 : B, T x0 y0
D:=(A * B)%type:Type
R:=fun z z' : D => let x0 := fst z in let x' := fst z' in let y0 := snd z in let y'0 := snd z' in x0 = x' /\ (T x0 y0 -> y0 = y'0 \/ T x0 y'0) /\ (T x0 y'0 -> y0 = y'0 \/ T x0 y0):D -> D -> Prop
Hequiv:Equivalence R
g:D -> D
T':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y'0 : B, T x0 y'0 /\ g (x0, y'0) = (x0, y0)):A -> B -> Prop
x:A
y:B
Hy:T x y
Heq1:fst (x, y) = fst (g (x, y))
H'':T (fst (x, y)) (snd (g (x, y))) -> snd (x, y) = snd (g (x, y)) \/ T (fst (x, y)) (snd (x, y))
Hgxyuniq:forall x' : D, R (x, y) x' -> g (x, y) = g x'
Hgy:T (fst (x, y)) (snd (g (x, y)))
y':B
Hy':T x y'
y'':B
Hy'':T x y''
Heq:g (x, y'') = (x, y')

(T (fst (x, y)) (snd (x, y)) -> snd (x, y) = snd (x, y'') \/ T (fst (x, y)) (snd (x, y''))) /\ (T (fst (x, y)) (snd (x, y'')) -> snd (x, y) = snd (x, y'') \/ T (fst (x, y)) (snd (x, y)))
split; right; easy. Qed. (**********************************************************************)

AC_fun_setoid = AC_fun_setoid_gen = AC_fun_setoid_simple


forall A B : Type, GeneralizedSetoidFunctionalChoice_on A B -> SetoidFunctionalChoice_on A B

forall A B : Type, GeneralizedSetoidFunctionalChoice_on A B -> SetoidFunctionalChoice_on A B
A, B:Type
GenSetoidFunChoice:GeneralizedSetoidFunctionalChoice_on A B
R:A -> A -> Prop
T:A -> B -> Prop
Hequiv:Equivalence R
Hcompat:forall (x x' : A) (y : B), R x x' -> T x y -> T x' y
Hex:forall x : A, exists y : B, T x y

exists f : A -> B, forall x : A, T x (f x) /\ (forall x' : A, R x x' -> f x = f x')
A, B:Type
GenSetoidFunChoice:GeneralizedSetoidFunctionalChoice_on A B
R:A -> A -> Prop
T:A -> B -> Prop
Hequiv:Equivalence R
Hcompat:forall (x x' : A) (y : B), R x x' -> T x y -> T x' y
Hex:forall x : A, exists y : B, T x y

Equivalence eq
A, B:Type
GenSetoidFunChoice:GeneralizedSetoidFunctionalChoice_on A B
R:A -> A -> Prop
T:A -> B -> Prop
Hequiv:Equivalence R
Hcompat:forall (x x' : A) (y : B), R x x' -> T x y -> T x' y
Hex:forall x : A, exists y : B, T x y
forall (x x' : A) (y y' : B), R x x' -> y = y' -> T x y -> T x' y'
A, B:Type
GenSetoidFunChoice:GeneralizedSetoidFunctionalChoice_on A B
R:A -> A -> Prop
T:A -> B -> Prop
Hequiv:Equivalence R
Hcompat:forall (x x' : A) (y : B), R x x' -> T x y -> T x' y
Hex:forall x : A, exists y : B, T x y

forall (x x' : A) (y y' : B), R x x' -> y = y' -> T x y -> T x' y'
A, B:Type
GenSetoidFunChoice:GeneralizedSetoidFunctionalChoice_on A B
R:A -> A -> Prop
T:A -> B -> Prop
Hequiv:Equivalence R
Hcompat:forall (x0 x'0 : A) (y0 : B), R x0 x'0 -> T x0 y0 -> T x'0 y0
Hex:forall x0 : A, exists y0 : B, T x0 y0
x, x':A
y:B
H:R x x'

T x y -> T x' y
firstorder. Qed.

forall A B : Type, SetoidFunctionalChoice_on A B -> GeneralizedSetoidFunctionalChoice_on A B

forall A B : Type, SetoidFunctionalChoice_on A B -> GeneralizedSetoidFunctionalChoice_on A B
A, B:Type
SetoidFunChoice:SetoidFunctionalChoice_on A B
R:A -> A -> Prop
S:B -> B -> Prop
T:A -> B -> Prop
HequivR:Equivalence R
HequivS:Equivalence S
Hcompat:forall (x x' : A) (y y' : B), R x x' -> S y y' -> T x y -> T x' y'
Hex:forall x : A, exists y : B, 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'))
A, B:Type
SetoidFunChoice:SetoidFunctionalChoice_on A B
R:A -> A -> Prop
S:B -> B -> Prop
T:A -> B -> Prop
HequivR:Equivalence R
HequivS:Equivalence S
Hcompat:forall (x x' : A) (y y' : B), R x x' -> S y y' -> T x y -> T x' y'
Hex:forall x : A, exists y : B, T x y

forall (x x' : A) (y : B), R x x' -> T x y -> T x' y
A, B:Type
SetoidFunChoice:SetoidFunctionalChoice_on A B
R:A -> A -> Prop
S:B -> B -> Prop
T:A -> B -> Prop
HequivR:Equivalence R
HequivS:Equivalence S
Hcompat:forall (x x' : A) (y y' : B), R x x' -> S y y' -> T x y -> T x' y'
Hex:forall x : A, exists y : B, T x y
f:A -> B
Hf:forall x : A, T x (f x) /\ (forall x' : A, R x x' -> f x = f x')
exists f0 : A -> B, forall x : A, T x (f0 x) /\ (forall x' : A, R x x' -> S (f0 x) (f0 x'))
A, B:Type
SetoidFunChoice:SetoidFunctionalChoice_on A B
R:A -> A -> Prop
S:B -> B -> Prop
T:A -> B -> Prop
HequivR:Equivalence R
HequivS:Equivalence S
Hcompat:forall (x x' : A) (y y' : B), R x x' -> S y y' -> T x y -> T x' y'
Hex:forall x : A, exists y : B, T x y

forall (x x' : A) (y : B), R x x' -> T x y -> T x' y
intros; apply (Hcompat x x' y y); try easy.
A, B:Type
SetoidFunChoice:SetoidFunctionalChoice_on A B
R:A -> A -> Prop
S:B -> B -> Prop
T:A -> B -> Prop
HequivR:Equivalence R
HequivS:Equivalence S
Hcompat:forall (x x' : A) (y y' : B), R x x' -> S y y' -> T x y -> T x' y'
Hex:forall x : A, exists y : B, T x y
f:A -> B
Hf:forall x : A, T x (f x) /\ (forall x' : A, R x x' -> f x = f x')

exists f0 : A -> B, forall x : A, T x (f0 x) /\ (forall x' : A, R x x' -> S (f0 x) (f0 x'))
A, B:Type
SetoidFunChoice:SetoidFunctionalChoice_on A B
R:A -> A -> Prop
S:B -> B -> Prop
T:A -> B -> Prop
HequivR:Equivalence R
HequivS:Equivalence S
Hcompat:forall (x x' : A) (y y' : B), R x x' -> S y y' -> T x y -> T x' y'
Hex:forall x : A, exists y : B, T x y
f:A -> B
Hf:forall x : A, T x (f x) /\ (forall x' : A, R x x' -> f x = f x')

forall x : A, T x (f x) /\ (forall x' : A, R x x' -> S (f x) (f x'))
A, B:Type
SetoidFunChoice:SetoidFunctionalChoice_on A B
R:A -> A -> Prop
S:B -> B -> Prop
T:A -> B -> Prop
HequivR:Equivalence R
HequivS:Equivalence S
Hcompat:forall (x0 x' : A) (y y' : B), R x0 x' -> S y y' -> T x0 y -> T x' y'
Hex:forall x0 : A, exists y : B, T x0 y
f:A -> B
x:A
Hf:T x (f x)
Huniq:forall x' : A, R x x' -> f x = f x'

T x (f x) /\ (forall x' : A, R x x' -> S (f x) (f x'))
A, B:Type
SetoidFunChoice:SetoidFunctionalChoice_on A B
R:A -> A -> Prop
S:B -> B -> Prop
T:A -> B -> Prop
HequivR:Equivalence R
HequivS:Equivalence S
Hcompat:forall (x0 x'0 : A) (y y' : B), R x0 x'0 -> S y y' -> T x0 y -> T x'0 y'
Hex:forall x0 : A, exists y : B, T x0 y
f:A -> B
x:A
Hf:T x (f x)
Huniq:forall x'0 : A, R x x'0 -> f x = f x'0
x':A
H:R x x'

S (f x) (f x')
now erewrite Huniq. Qed.

forall A B : Type, SetoidFunctionalChoice_on A B <-> GeneralizedSetoidFunctionalChoice_on A B

forall A B : Type, SetoidFunctionalChoice_on A B <-> GeneralizedSetoidFunctionalChoice_on A B
split; auto using gen_setoid_fun_choice_imp_setoid_fun_choice, setoid_fun_choice_imp_gen_setoid_fun_choice. Qed.

forall A B : Type, SetoidFunctionalChoice_on A B -> SimpleSetoidFunctionalChoice_on A B

forall A B : Type, SetoidFunctionalChoice_on A B -> SimpleSetoidFunctionalChoice_on A B
A, B:Type
SetoidFunChoice:SetoidFunctionalChoice_on A B
R:A -> A -> Prop
T:A -> B -> Prop
Hequiv:Equivalence R
Hexists:forall x : A, exists y : B, forall x' : A, 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')
A, B:Type
SetoidFunChoice:SetoidFunctionalChoice_on A B
R:A -> A -> Prop
T:A -> B -> Prop
Hequiv:Equivalence R
Hexists:forall x : A, exists y : B, forall x' : A, R x x' -> T x' y
T':=fun (x : A) (y : B) => forall x' : A, R x x' -> T x' y:A -> B -> Prop

exists f : A -> B, forall x : A, T x (f x) /\ (forall x' : A, R x x' -> f x = f x')
A, B:Type
SetoidFunChoice:SetoidFunctionalChoice_on A B
R:A -> A -> Prop
T:A -> B -> Prop
Hequiv:Equivalence R
Hexists:forall x : A, exists y : B, forall x' : A, R x x' -> T x' y
T':=fun (x : A) (y : B) => forall x' : A, R x x' -> T x' y:A -> B -> Prop
Hcompat:forall (x x' : A) (y : B), R x x' -> T' x 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')
A, B:Type
SetoidFunChoice:SetoidFunctionalChoice_on A B
R:A -> A -> Prop
T:A -> B -> Prop
Hequiv:Equivalence R
Hexists:forall x : A, exists y : B, forall x' : A, R x x' -> T x' y
T':=fun (x : A) (y : B) => forall x' : A, R x x' -> T x' y:A -> B -> Prop
Hcompat:forall (x x' : A) (y : B), R x x' -> T' x y -> T' x' y
f:A -> B
Hf:forall x : A, T' x (f x) /\ (forall x' : A, R x x' -> f x = f x')

exists f0 : A -> B, forall x : A, T x (f0 x) /\ (forall x' : A, R x x' -> f0 x = f0 x')
A, B:Type
SetoidFunChoice:SetoidFunctionalChoice_on A B
R:A -> A -> Prop
T:A -> B -> Prop
Hequiv:Equivalence R
Hexists:forall x : A, exists y : B, forall x' : A, R x x' -> T x' y
T':=fun (x : A) (y : B) => forall x' : A, R x x' -> T x' y:A -> B -> Prop
Hcompat:forall (x x' : A) (y : B), R x x' -> T' x y -> T' x' y
f:A -> B
Hf:forall x : A, T' x (f x) /\ (forall x' : A, R x x' -> f x = f x')

forall x : A, T x (f x) /\ (forall x' : A, R x x' -> f x = f x')
firstorder. Qed.

forall A B : Type, SimpleSetoidFunctionalChoice_on A B -> SetoidFunctionalChoice_on A B

forall A B : Type, SimpleSetoidFunctionalChoice_on A B -> SetoidFunctionalChoice_on A B
A, B:Type
SimpleSetoidFunChoice:SimpleSetoidFunctionalChoice_on A B
R:A -> A -> Prop
T:A -> B -> Prop
Hequiv:Equivalence R
Hcompat:forall (x x' : A) (y : B), R x x' -> T x y -> T x' y
Hexists:forall x : A, exists y : B, T x y

exists f : A -> B, forall x : A, T x (f x) /\ (forall x' : A, R x x' -> f x = f x')
destruct (SimpleSetoidFunChoice R T Hequiv) as (f,Hf); firstorder. Qed.

forall A B : Type, SetoidFunctionalChoice_on A B <-> SimpleSetoidFunctionalChoice_on A B

forall A B : Type, SetoidFunctionalChoice_on A B <-> SimpleSetoidFunctionalChoice_on A B
split; auto using simple_setoid_fun_choice_imp_setoid_fun_choice, setoid_fun_choice_imp_simple_setoid_fun_choice. Qed. (**********************************************************************)

AC_fun_setoid = AC! + AC_fun_repr


forall A B : Type, SetoidFunctionalChoice_on A B -> FunctionalChoice_on A B

forall A B : Type, SetoidFunctionalChoice_on A B -> FunctionalChoice_on A B
A, B:Type
SetoidFunChoice:SetoidFunctionalChoice_on A B
T:A -> B -> Prop
Hexists:forall x : A, exists y : B, T x y

exists f : A -> B, forall x : A, T x (f x)
A, B:Type
SetoidFunChoice:SetoidFunctionalChoice_on A B
T:A -> B -> Prop
Hexists:forall x : A, exists y : B, T x y

Equivalence eq
A, B:Type
SetoidFunChoice:SetoidFunctionalChoice_on A B
T:A -> B -> Prop
Hexists:forall x : A, exists y : B, T x y
forall (x x' : A) (y : B), x = x' -> T x y -> T x' y
A, B:Type
SetoidFunChoice:SetoidFunctionalChoice_on A B
T:A -> B -> Prop
Hexists:forall x : A, exists y : B, T x y
forall x : A, exists y : B, T x y
A, B:Type
SetoidFunChoice:SetoidFunctionalChoice_on A B
T:A -> B -> Prop
Hexists:forall x : A, exists y : B, T x y
f:A -> B
Hf:forall x : A, T x (f x) /\ (forall x' : A, x = x' -> f x = f x')
exists f0 : A -> B, forall x : A, T x (f0 x)
A, B:Type
SetoidFunChoice:SetoidFunctionalChoice_on A B
T:A -> B -> Prop
Hexists:forall x : A, exists y : B, T x y

Equivalence eq
apply eq_equivalence.
A, B:Type
SetoidFunChoice:SetoidFunctionalChoice_on A B
T:A -> B -> Prop
Hexists:forall x : A, exists y : B, T x y

forall (x x' : A) (y : B), x = x' -> T x y -> T x' y
now intros * ->.
A, B:Type
SetoidFunChoice:SetoidFunctionalChoice_on A B
T:A -> B -> Prop
Hexists:forall x : A, exists y : B, T x y

forall x : A, exists y : B, T x y
assumption.
A, B:Type
SetoidFunChoice:SetoidFunctionalChoice_on A B
T:A -> B -> Prop
Hexists:forall x : A, exists y : B, T x y
f:A -> B
Hf:forall x : A, T x (f x) /\ (forall x' : A, x = x' -> f x = f x')

exists f0 : A -> B, forall x : A, T x (f0 x)
A, B:Type
SetoidFunChoice:SetoidFunctionalChoice_on A B
T:A -> B -> Prop
Hexists:forall x : A, exists y : B, T x y
f:A -> B
Hf:forall x : A, T x (f x) /\ (forall x' : A, x = x' -> f x = f x')

forall x : A, T x (f x)
firstorder. Qed.

forall A B : Type, SetoidFunctionalChoice_on A B -> FunctionalRelReification_on A B

forall A B : Type, SetoidFunctionalChoice_on A B -> FunctionalRelReification_on A B
A, B:Type
SetoidFunChoice:SetoidFunctionalChoice_on A B

FunctionalRelReification_on A B
A, B:Type
SetoidFunChoice:SetoidFunctionalChoice_on A B

FunctionalChoice_on A B
now apply setoid_fun_choice_imp_fun_choice. Qed.

SetoidFunctionalChoice -> RepresentativeFunctionalChoice

SetoidFunctionalChoice -> RepresentativeFunctionalChoice
SetoidFunChoice:SetoidFunctionalChoice
A:Type
R:A -> A -> Prop
Hequiv:Equivalence R

exists f : A -> A, forall x : A, R x (f x) /\ (forall x' : A, R x x' -> f x = f x')
apply SetoidFunChoice; firstorder. Qed.

FunctionalRelReification -> RepresentativeFunctionalChoice -> SetoidFunctionalChoice

FunctionalRelReification -> RepresentativeFunctionalChoice -> SetoidFunctionalChoice
FunRelReify:FunctionalRelReification
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
R:A -> A -> Prop
T:A -> B -> Prop
Hequiv:Equivalence R
Hcompat:forall (x x' : A) (y : B), R x x' -> T x y -> T x' y
Hexists:forall x : A, exists y : B, T x y

exists f : A -> B, forall x : A, T x (f x) /\ (forall x' : A, R x x' -> f x = f x')
FunRelReify:FunctionalRelReification
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
R:A -> A -> Prop
T:A -> B -> Prop
Hequiv:Equivalence R
Hcompat:forall (x x' : A) (y : B), R x x' -> T x y -> T x' y
Hexists:forall x : A, exists y : B, T x y

FunctionalChoice
FunRelReify:FunctionalRelReification
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
R:A -> A -> Prop
T:A -> B -> Prop
Hequiv:Equivalence R
Hcompat:forall (x x' : A) (y : B), R x x' -> T x y -> T x' y
Hexists:forall x : A, exists y : B, T x y
FunChoice:FunctionalChoice
exists f : A -> B, forall x : A, T x (f x) /\ (forall x' : A, R x x' -> f x = f x')
FunRelReify:FunctionalRelReification
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
R:A -> A -> Prop
T:A -> B -> Prop
Hequiv:Equivalence R
Hcompat:forall (x x' : A) (y : B), R x x' -> T x y -> T x' y
Hexists:forall x : A, exists y : B, T x y

FunctionalChoice
FunRelReify:FunctionalRelReification
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
R:A -> A -> Prop
T:A -> B -> Prop
Hequiv:Equivalence R
Hcompat:forall (x x' : A) (y : B), R x x' -> T x y -> T x' y
Hexists:forall x : A, exists y : B, T x y
A', B':Type

FunctionalChoice_on A' B'
FunRelReify:FunctionalRelReification
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
R:A -> A -> Prop
T:A -> B -> Prop
Hequiv:Equivalence R
Hcompat:forall (x x' : A) (y : B), R x x' -> T x y -> T x' y
Hexists:forall x : A, exists y : B, T x y
A', B':Type

FunctionalRelReification_on A' B'
FunRelReify:FunctionalRelReification
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
R:A -> A -> Prop
T:A -> B -> Prop
Hequiv:Equivalence R
Hcompat:forall (x x' : A) (y : B), R x x' -> T x y -> T x' y
Hexists:forall x : A, exists y : B, T x y
A', B':Type
RelationalChoice_on A' B'
FunRelReify:FunctionalRelReification
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
R:A -> A -> Prop
T:A -> B -> Prop
Hequiv:Equivalence R
Hcompat:forall (x x' : A) (y : B), R x x' -> T x y -> T x' y
Hexists:forall x : A, exists y : B, T x y
A', B':Type

FunctionalRelReification_on A' B'
apply FunRelReify.
FunRelReify:FunctionalRelReification
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
R:A -> A -> Prop
T:A -> B -> Prop
Hequiv:Equivalence R
Hcompat:forall (x x' : A) (y : B), R x x' -> T x y -> T x' y
Hexists:forall x : A, exists y : B, T x y
A', B':Type

RelationalChoice_on A' B'
now apply repr_fun_choice_imp_relational_choice.
FunRelReify:FunctionalRelReification
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
R:A -> A -> Prop
T:A -> B -> Prop
Hequiv:Equivalence R
Hcompat:forall (x x' : A) (y : B), R x x' -> T x y -> T x' y
Hexists:forall x : A, exists y : B, T x y
FunChoice:FunctionalChoice

exists f : A -> B, forall x : A, T x (f x) /\ (forall x' : A, R x x' -> f x = f x')
FunRelReify:FunctionalRelReification
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
R:A -> A -> Prop
T:A -> B -> Prop
Hequiv:Equivalence R
Hcompat:forall (x x' : A) (y : B), R x x' -> T x y -> T x' y
Hexists:forall x : A, exists y : B, T x y
FunChoice:FunctionalChoice
f:A -> B
Hf:forall x : A, T x (f x)

exists f0 : A -> B, forall x : A, T x (f0 x) /\ (forall x' : A, R x x' -> f0 x = f0 x')
FunRelReify:FunctionalRelReification
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
R:A -> A -> Prop
T:A -> B -> Prop
Hequiv:Equivalence R
Hcompat:forall (x x' : A) (y : B), R x x' -> T x y -> T x' y
Hexists:forall x : A, exists y : B, T x y
FunChoice:FunctionalChoice
f:A -> B
Hf:forall x : A, T x (f x)
g:A -> A
Hg:forall x : A, R x (g x) /\ (forall x' : A, R x x' -> g x = g x')

exists f0 : A -> B, forall x : A, T x (f0 x) /\ (forall x' : A, R x x' -> f0 x = f0 x')
FunRelReify:FunctionalRelReification
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
R:A -> A -> Prop
T:A -> B -> Prop
Hequiv:Equivalence R
Hcompat:forall (x x' : A) (y : B), R x x' -> T x y -> T x' y
Hexists:forall x : A, exists y : B, T x y
FunChoice:FunctionalChoice
f:A -> B
Hf:forall x : A, T x (f x)
g:A -> A
Hg:forall x : A, R x (g x) /\ (forall x' : A, R x x' -> g x = g x')

forall x : A, T x (f (g x)) /\ (forall x' : A, R x x' -> f (g x) = f (g x'))
FunRelReify:FunctionalRelReification
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
R:A -> A -> Prop
T:A -> B -> Prop
Hequiv:Equivalence R
Hcompat:forall (x0 x' : A) (y : B), R x0 x' -> T x0 y -> T x' y
Hexists:forall x0 : A, exists y : B, T x0 y
FunChoice:FunctionalChoice
f:A -> B
Hf:forall x0 : A, T x0 (f x0)
g:A -> A
Hg:forall x0 : A, R x0 (g x0) /\ (forall x' : A, R x0 x' -> g x0 = g x')
x:A

T x (f (g x)) /\ (forall x' : A, R x x' -> f (g x) = f (g x'))
FunRelReify:FunctionalRelReification
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
R:A -> A -> Prop
T:A -> B -> Prop
Hequiv:Equivalence R
Hcompat:forall (x0 x' : A) (y : B), R x0 x' -> T x0 y -> T x' y
Hexists:forall x0 : A, exists y : B, T x0 y
FunChoice:FunctionalChoice
f:A -> B
Hf:forall x0 : A, T x0 (f x0)
g:A -> A
Hg:forall x0 : A, R x0 (g x0) /\ (forall x' : A, R x0 x' -> g x0 = g x')
x:A
Hgx:R x (g x)
HRuniq:forall x' : A, R x x' -> g x = g x'

T x (f (g x)) /\ (forall x' : A, R x x' -> f (g x) = f (g x'))
FunRelReify:FunctionalRelReification
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
R:A -> A -> Prop
T:A -> B -> Prop
Hequiv:Equivalence R
Hcompat:forall (x0 x' : A) (y : B), R x0 x' -> T x0 y -> T x' y
Hexists:forall x0 : A, exists y : B, T x0 y
FunChoice:FunctionalChoice
f:A -> B
Hf:forall x0 : A, T x0 (f x0)
g:A -> A
Hg:forall x0 : A, R x0 (g x0) /\ (forall x' : A, R x0 x' -> g x0 = g x')
x:A
Hgx:R x (g x)
HRuniq:forall x' : A, R x x' -> g x = g x'

T x (f (g x))
FunRelReify:FunctionalRelReification
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
R:A -> A -> Prop
T:A -> B -> Prop
Hequiv:Equivalence R
Hcompat:forall (x0 x' : A) (y : B), R x0 x' -> T x0 y -> T x' y
Hexists:forall x0 : A, exists y : B, T x0 y
FunChoice:FunctionalChoice
f:A -> B
Hf:forall x0 : A, T x0 (f x0)
g:A -> A
Hg:forall x0 : A, R x0 (g x0) /\ (forall x' : A, R x0 x' -> g x0 = g x')
x:A
Hgx:R x (g x)
HRuniq:forall x' : A, R x x' -> g x = g x'
forall x' : A, R x x' -> f (g x) = f (g x')
FunRelReify:FunctionalRelReification
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
R:A -> A -> Prop
T:A -> B -> Prop
Hequiv:Equivalence R
Hcompat:forall (x0 x' : A) (y : B), R x0 x' -> T x0 y -> T x' y
Hexists:forall x0 : A, exists y : B, T x0 y
FunChoice:FunctionalChoice
f:A -> B
Hf:forall x0 : A, T x0 (f x0)
g:A -> A
Hg:forall x0 : A, R x0 (g x0) /\ (forall x' : A, R x0 x' -> g x0 = g x')
x:A
Hgx:R x (g x)
HRuniq:forall x' : A, R x x' -> g x = g x'

T x (f (g x))
FunRelReify:FunctionalRelReification
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
R:A -> A -> Prop
T:A -> B -> Prop
Hequiv:Equivalence R
Hcompat:forall (x0 x' : A) (y : B), R x0 x' -> T x0 y -> T x' y
Hexists:forall x0 : A, exists y : B, T x0 y
FunChoice:FunctionalChoice
f:A -> B
Hf:forall x0 : A, T x0 (f x0)
g:A -> A
Hg:forall x0 : A, R x0 (g x0) /\ (forall x' : A, R x0 x' -> g x0 = g x')
x:A
Hgx:R x (g x)
HRuniq:forall x' : A, R x x' -> g x = g x'

R ?x x
FunRelReify:FunctionalRelReification
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
R:A -> A -> Prop
T:A -> B -> Prop
Hequiv:Equivalence R
Hcompat:forall (x0 x' : A) (y : B), R x0 x' -> T x0 y -> T x' y
Hexists:forall x0 : A, exists y : B, T x0 y
FunChoice:FunctionalChoice
f:A -> B
Hf:forall x0 : A, T x0 (f x0)
g:A -> A
Hg:forall x0 : A, R x0 (g x0) /\ (forall x' : A, R x0 x' -> g x0 = g x')
x:A
Hgx:R x (g x)
HRuniq:forall x' : A, R x x' -> g x = g x'
T ?x (f (g x))
FunRelReify:FunctionalRelReification
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
R:A -> A -> Prop
T:A -> B -> Prop
Hequiv:Equivalence R
Hcompat:forall (x0 x' : A) (y : B), R x0 x' -> T x0 y -> T x' y
Hexists:forall x0 : A, exists y : B, T x0 y
FunChoice:FunctionalChoice
f:A -> B
Hf:forall x0 : A, T x0 (f x0)
g:A -> A
Hg:forall x0 : A, R x0 (g x0) /\ (forall x' : A, R x0 x' -> g x0 = g x')
x:A
Hgx:R x (g x)
HRuniq:forall x' : A, R x x' -> g x = g x'

R x ?x
FunRelReify:FunctionalRelReification
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
R:A -> A -> Prop
T:A -> B -> Prop
Hequiv:Equivalence R
Hcompat:forall (x0 x' : A) (y : B), R x0 x' -> T x0 y -> T x' y
Hexists:forall x0 : A, exists y : B, T x0 y
FunChoice:FunctionalChoice
f:A -> B
Hf:forall x0 : A, T x0 (f x0)
g:A -> A
Hg:forall x0 : A, R x0 (g x0) /\ (forall x' : A, R x0 x' -> g x0 = g x')
x:A
Hgx:R x (g x)
HRuniq:forall x' : A, R x x' -> g x = g x'
T ?x (f (g x))
FunRelReify:FunctionalRelReification
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
R:A -> A -> Prop
T:A -> B -> Prop
Hequiv:Equivalence R
Hcompat:forall (x0 x' : A) (y : B), R x0 x' -> T x0 y -> T x' y
Hexists:forall x0 : A, exists y : B, T x0 y
FunChoice:FunctionalChoice
f:A -> B
Hf:forall x0 : A, T x0 (f x0)
g:A -> A
Hg:forall x0 : A, R x0 (g x0) /\ (forall x' : A, R x0 x' -> g x0 = g x')
x:A
Hgx:R x (g x)
HRuniq:forall x' : A, R x x' -> g x = g x'

T (g x) (f (g x))
apply Hf.
FunRelReify:FunctionalRelReification
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
R:A -> A -> Prop
T:A -> B -> Prop
Hequiv:Equivalence R
Hcompat:forall (x0 x' : A) (y : B), R x0 x' -> T x0 y -> T x' y
Hexists:forall x0 : A, exists y : B, T x0 y
FunChoice:FunctionalChoice
f:A -> B
Hf:forall x0 : A, T x0 (f x0)
g:A -> A
Hg:forall x0 : A, R x0 (g x0) /\ (forall x' : A, R x0 x' -> g x0 = g x')
x:A
Hgx:R x (g x)
HRuniq:forall x' : A, R x x' -> g x = g x'

forall x' : A, R x x' -> f (g x) = f (g x')
FunRelReify:FunctionalRelReification
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
R:A -> A -> Prop
T:A -> B -> Prop
Hequiv:Equivalence R
Hcompat:forall (x0 x' : A) (y0 : B), R x0 x' -> T x0 y0 -> T x' y0
Hexists:forall x0 : A, exists y0 : B, T x0 y0
FunChoice:FunctionalChoice
f:A -> B
Hf:forall x0 : A, T x0 (f x0)
g:A -> A
Hg:forall x0 : A, R x0 (g x0) /\ (forall x' : A, R x0 x' -> g x0 = g x')
x:A
Hgx:R x (g x)
HRuniq:forall x' : A, R x x' -> g x = g x'
y:A
Hxy:R x y

f (g x) = f (g y)
FunRelReify:FunctionalRelReification
ReprFunChoice:RepresentativeFunctionalChoice
A, B:Type
R:A -> A -> Prop
T:A -> B -> Prop
Hequiv:Equivalence R
Hcompat:forall (x0 x' : A) (y0 : B), R x0 x' -> T x0 y0 -> T x' y0
Hexists:forall x0 : A, exists y0 : B, T x0 y0
FunChoice:FunctionalChoice
f:A -> B
Hf:forall x0 : A, T x0 (f x0)
g:A -> A
Hg:forall x0 : A, R x0 (g x0) /\ (forall x' : A, R x0 x' -> g x0 = g x')
x:A
Hgx:R x (g x)
HRuniq:forall x' : A, R x x' -> g x = g x'
y:A
Hxy:R x y

g x = g y
auto. Qed.

FunctionalRelReification /\ RepresentativeFunctionalChoice <-> SetoidFunctionalChoice

FunctionalRelReification /\ RepresentativeFunctionalChoice <-> SetoidFunctionalChoice
H:FunctionalRelReification /\ RepresentativeFunctionalChoice
A, B:Type

SetoidFunctionalChoice_on A B
H:SetoidFunctionalChoice
FunctionalRelReification /\ RepresentativeFunctionalChoice
H:FunctionalRelReification /\ RepresentativeFunctionalChoice
A, B:Type

SetoidFunctionalChoice_on A B
now apply functional_rel_reification_and_repr_fun_choice_imp_setoid_fun_choice.
H:SetoidFunctionalChoice

FunctionalRelReification /\ RepresentativeFunctionalChoice
H:SetoidFunctionalChoice

FunctionalRelReification
H:SetoidFunctionalChoice
RepresentativeFunctionalChoice
H:SetoidFunctionalChoice

FunctionalRelReification
now intros A B; apply setoid_fun_choice_imp_functional_rel_reification.
H:SetoidFunctionalChoice

RepresentativeFunctionalChoice
now apply setoid_fun_choice_imp_repr_fun_choice. Qed.
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

FunctionalChoice -> ExtensionalFunctionRepresentative -> ExcludedMiddle -> RepresentativeFunctionalChoice

FunctionalChoice -> ExtensionalFunctionRepresentative -> ExcludedMiddle -> RepresentativeFunctionalChoice
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R

exists f : A -> A, forall x : A, R x (f x) /\ (forall x' : A, R x x' -> f x = f x')
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R

forall P : Prop, exists b : bool, b = true <-> P
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P
exists f : A -> A, forall x : A, R x (f x) /\ (forall x' : A, R x x' -> f x = f x')
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R

forall P : Prop, exists b : bool, b = true <-> P
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
P:Prop

exists b : bool, b = true <-> P
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
P:Prop
H:P

exists b : bool, b = true <-> P
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
P:Prop
H:~ P
exists b : bool, b = true <-> P
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
P:Prop
H:P

exists b : bool, b = true <-> P
exists true; firstorder.
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
P:Prop
H:~ P

exists b : bool, b = true <-> P
exists false; easy.
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P

exists f : A -> A, forall x : A, R x (f x) /\ (forall x' : A, R x x' -> f x = f x')
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P
c:Prop -> bool
Hc:forall x : Prop, c x = true <-> x

exists f : A -> A, forall x : A, R x (f x) /\ (forall x' : A, R x x' -> f x = f x')
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P
c:Prop -> bool
Hc:forall x : Prop, c x = true <-> x
class_of:=fun a y : A => c (R a y):A -> A -> bool

exists f : A -> A, forall x : A, R x (f x) /\ (forall x' : A, R x x' -> f x = f x')
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P
c:Prop -> bool
Hc:forall x : Prop, c x = true <-> x
class_of:=fun a y : A => c (R a y):A -> A -> bool
isclass:=fun f : A -> bool => exists x : A, f x = true:(A -> bool) -> Prop

exists f : A -> A, forall x : A, R x (f x) /\ (forall x' : A, R x x' -> f x = f x')
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P
c:Prop -> bool
Hc:forall x : Prop, c x = true <-> x
class_of:=fun a y : A => c (R a y):A -> A -> bool
isclass:=fun f : A -> bool => exists x : A, f x = true:(A -> bool) -> Prop
class:={f : A -> bool | isclass f}:Type

exists f : A -> A, forall x : A, R x (f x) /\ (forall x' : A, R x x' -> f x = f x')
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P
c:Prop -> bool
Hc:forall x : Prop, c x = true <-> x
class_of:=fun a y : A => c (R a y):A -> A -> bool
isclass:=fun f : A -> bool => exists x : A, f x = true:(A -> bool) -> Prop
class:={f : A -> bool | isclass f}:Type
contains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Prop

exists f : A -> A, forall x : A, R x (f x) /\ (forall x' : A, R x x' -> f x = f x')
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P
c:Prop -> bool
Hc:forall x : Prop, c x = true <-> x
class_of:=fun a y : A => c (R a y):A -> A -> bool
isclass:=fun f : A -> bool => exists x : A, f x = true:(A -> bool) -> Prop
class:={f : A -> bool | isclass f}:Type
contains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Prop

forall x : class, exists y : A, contains x y
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P
c:Prop -> bool
Hc:forall x : Prop, c x = true <-> x
class_of:=fun a y : A => c (R a y):A -> A -> bool
isclass:=fun f0 : A -> bool => exists x : A, f0 x = true:(A -> bool) -> Prop
class:={f0 : A -> bool | isclass f0}:Type
contains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Prop
f:class -> A
Hf:forall x : class, contains x (f x)
exists f0 : A -> A, forall x : A, R x (f0 x) /\ (forall x' : A, R x x' -> f0 x = f0 x')
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P
c:Prop -> bool
Hc:forall x : Prop, c x = true <-> x
class_of:=fun a y : A => c (R a y):A -> A -> bool
isclass:=fun f : A -> bool => exists x : A, f x = true:(A -> bool) -> Prop
class:={f : A -> bool | isclass f}:Type
contains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Prop

forall x : class, exists y : A, contains x y
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P
c:Prop -> bool
Hc:forall x : Prop, c x = true <-> x
class_of:=fun a y : A => c (R a y):A -> A -> bool
isclass:=fun f0 : A -> bool => exists x : A, f0 x = true:(A -> bool) -> Prop
class:={f0 : A -> bool | isclass f0}:Type
contains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Prop
f:class

exists y : A, contains f y
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P
c:Prop -> bool
Hc:forall x0 : Prop, c x0 = true <-> x0
class_of:=fun a y : A => c (R a y):A -> A -> bool
isclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Prop
class:={f0 : A -> bool | isclass f0}:Type
contains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Prop
f:class
x:A
Hx:proj1_sig f x = true

exists y : A, contains f y
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P
c:Prop -> bool
Hc:forall x0 : Prop, c x0 = true <-> x0
class_of:=fun a y : A => c (R a y):A -> A -> bool
isclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Prop
class:={f0 : A -> bool | isclass f0}:Type
contains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Prop
f:class
x:A
Hx:proj1_sig f x = true

contains f x
easy.
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P
c:Prop -> bool
Hc:forall x : Prop, c x = true <-> x
class_of:=fun a y : A => c (R a y):A -> A -> bool
isclass:=fun f0 : A -> bool => exists x : A, f0 x = true:(A -> bool) -> Prop
class:={f0 : A -> bool | isclass f0}:Type
contains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Prop
f:class -> A
Hf:forall x : class, contains x (f x)

exists f0 : A -> A, forall x : A, R x (f0 x) /\ (forall x' : A, R x x' -> f0 x = f0 x')
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P
c:Prop -> bool
Hc:forall x : Prop, c x = true <-> x
class_of:=fun a y : A => c (R a y):A -> A -> bool
isclass:=fun f0 : A -> bool => exists x : A, f0 x = true:(A -> bool) -> Prop
class:={f0 : A -> bool | isclass f0}:Type
contains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Prop
f:class -> A
Hf:forall x : class, contains x (f x)
h:(A -> bool) -> A -> bool
Hh:forall f0 : A -> bool, (forall x : A, f0 x = h f0 x) /\ (forall g : A -> bool, (forall x : A, f0 x = g x) -> h f0 = h g)

exists f0 : A -> A, forall x : A, R x (f0 x) /\ (forall x' : A, R x x' -> f0 x = f0 x')
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P
c:Prop -> bool
Hc:forall x : Prop, c x = true <-> x
class_of:=fun a y : A => c (R a y):A -> A -> bool
isclass:=fun f0 : A -> bool => exists x : A, f0 x = true:(A -> bool) -> Prop
class:={f0 : A -> bool | isclass f0}:Type
contains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Prop
f:class -> A
Hf:forall x : class, contains x (f x)
h:(A -> bool) -> A -> bool
Hh:forall f0 : A -> bool, (forall x : A, f0 x = h f0 x) /\ (forall g : A -> bool, (forall x : A, f0 x = g x) -> h f0 = h g)

forall a : A, isclass (h (class_of a))
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P
c:Prop -> bool
Hc:forall x : Prop, c x = true <-> x
class_of:=fun a y : A => c (R a y):A -> A -> bool
isclass:=fun f0 : A -> bool => exists x : A, f0 x = true:(A -> bool) -> Prop
class:={f0 : A -> bool | isclass f0}:Type
contains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Prop
f:class -> A
Hf:forall x : class, contains x (f x)
h:(A -> bool) -> A -> bool
Hh:forall f0 : A -> bool, (forall x : A, f0 x = h f0 x) /\ (forall g : A -> bool, (forall x : A, f0 x = g x) -> h f0 = h g)
Hisclass:forall a : A, isclass (h (class_of a))
exists f0 : A -> A, forall x : A, R x (f0 x) /\ (forall x' : A, R x x' -> f0 x = f0 x')
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P
c:Prop -> bool
Hc:forall x : Prop, c x = true <-> x
class_of:=fun a y : A => c (R a y):A -> A -> bool
isclass:=fun f0 : A -> bool => exists x : A, f0 x = true:(A -> bool) -> Prop
class:={f0 : A -> bool | isclass f0}:Type
contains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Prop
f:class -> A
Hf:forall x : class, contains x (f x)
h:(A -> bool) -> A -> bool
Hh:forall f0 : A -> bool, (forall x : A, f0 x = h f0 x) /\ (forall g : A -> bool, (forall x : A, f0 x = g x) -> h f0 = h g)

forall a : A, isclass (h (class_of a))
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P
c:Prop -> bool
Hc:forall x : Prop, c x = true <-> x
class_of:=fun a0 y : A => c (R a0 y):A -> A -> bool
isclass:=fun f0 : A -> bool => exists x : A, f0 x = true:(A -> bool) -> Prop
class:={f0 : A -> bool | isclass f0}:Type
contains:=fun (c0 : class) (a0 : A) => proj1_sig c0 a0 = true:class -> A -> Prop
f:class -> A
Hf:forall x : class, contains x (f x)
h:(A -> bool) -> A -> bool
Hh:forall f0 : A -> bool, (forall x : A, f0 x = h f0 x) /\ (forall g : A -> bool, (forall x : A, f0 x = g x) -> h f0 = h g)
a:A

isclass (h (class_of a))
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P
c:Prop -> bool
Hc:forall x : Prop, c x = true <-> x
class_of:=fun a0 y : A => c (R a0 y):A -> A -> bool
isclass:=fun f0 : A -> bool => exists x : A, f0 x = true:(A -> bool) -> Prop
class:={f0 : A -> bool | isclass f0}:Type
contains:=fun (c0 : class) (a0 : A) => proj1_sig c0 a0 = true:class -> A -> Prop
f:class -> A
Hf:forall x : class, contains x (f x)
h:(A -> bool) -> A -> bool
Hh:forall f0 : A -> bool, (forall x : A, f0 x = h f0 x) /\ (forall g : A -> bool, (forall x : A, f0 x = g x) -> h f0 = h g)
a:A

h (class_of a) a = true
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P
c:Prop -> bool
Hc:forall x : Prop, c x = true <-> x
class_of:=fun a0 y : A => c (R a0 y):A -> A -> bool
isclass:=fun f0 : A -> bool => exists x : A, f0 x = true:(A -> bool) -> Prop
class:={f0 : A -> bool | isclass f0}:Type
contains:=fun (c0 : class) (a0 : A) => proj1_sig c0 a0 = true:class -> A -> Prop
f:class -> A
Hf:forall x : class, contains x (f x)
h:(A -> bool) -> A -> bool
Hh:forall f0 : A -> bool, (forall x : A, f0 x = h f0 x) /\ (forall g : A -> bool, (forall x : A, f0 x = g x) -> h f0 = h g)
a:A
Ha:forall x : A, class_of a x = h (class_of a) x
Huniqa:forall g : A -> bool, (forall x : A, class_of a x = g x) -> h (class_of a) = h g

h (class_of a) a = true
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P
c:Prop -> bool
Hc:forall x : Prop, c x = true <-> x
class_of:=fun a0 y : A => c (R a0 y):A -> A -> bool
isclass:=fun f0 : A -> bool => exists x : A, f0 x = true:(A -> bool) -> Prop
class:={f0 : A -> bool | isclass f0}:Type
contains:=fun (c0 : class) (a0 : A) => proj1_sig c0 a0 = true:class -> A -> Prop
f:class -> A
Hf:forall x : class, contains x (f x)
h:(A -> bool) -> A -> bool
Hh:forall f0 : A -> bool, (forall x : A, f0 x = h f0 x) /\ (forall g : A -> bool, (forall x : A, f0 x = g x) -> h f0 = h g)
a:A
Ha:forall x : A, class_of a x = h (class_of a) x
Huniqa:forall g : A -> bool, (forall x : A, class_of a x = g x) -> h (class_of a) = h g

class_of a a = true
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P
c:Prop -> bool
Hc:forall x : Prop, c x = true <-> x
class_of:=fun a0 y : A => c (R a0 y):A -> A -> bool
isclass:=fun f0 : A -> bool => exists x : A, f0 x = true:(A -> bool) -> Prop
class:={f0 : A -> bool | isclass f0}:Type
contains:=fun (c0 : class) (a0 : A) => proj1_sig c0 a0 = true:class -> A -> Prop
f:class -> A
Hf:forall x : class, contains x (f x)
h:(A -> bool) -> A -> bool
Hh:forall f0 : A -> bool, (forall x : A, f0 x = h f0 x) /\ (forall g : A -> bool, (forall x : A, f0 x = g x) -> h f0 = h g)
a:A
Ha:forall x : A, class_of a x = h (class_of a) x
Huniqa:forall g : A -> bool, (forall x : A, class_of a x = g x) -> h (class_of a) = h g

R a a
apply Hrefl.
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P
c:Prop -> bool
Hc:forall x : Prop, c x = true <-> x
class_of:=fun a y : A => c (R a y):A -> A -> bool
isclass:=fun f0 : A -> bool => exists x : A, f0 x = true:(A -> bool) -> Prop
class:={f0 : A -> bool | isclass f0}:Type
contains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Prop
f:class -> A
Hf:forall x : class, contains x (f x)
h:(A -> bool) -> A -> bool
Hh:forall f0 : A -> bool, (forall x : A, f0 x = h f0 x) /\ (forall g : A -> bool, (forall x : A, f0 x = g x) -> h f0 = h g)
Hisclass:forall a : A, isclass (h (class_of a))

exists f0 : A -> A, forall x : A, R x (f0 x) /\ (forall x' : A, R x x' -> f0 x = f0 x')
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P
c:Prop -> bool
Hc:forall x : Prop, c x = true <-> x
class_of:=fun a y : A => c (R a y):A -> A -> bool
isclass:=fun f0 : A -> bool => exists x : A, f0 x = true:(A -> bool) -> Prop
class:={f0 : A -> bool | isclass f0}:Type
contains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Prop
f:class -> A
Hf:forall x : class, contains x (f x)
h:(A -> bool) -> A -> bool
Hh:forall f0 : A -> bool, (forall x : A, f0 x = h f0 x) /\ (forall g : A -> bool, (forall x : A, f0 x = g x) -> h f0 = h g)
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class

exists f0 : A -> A, forall x : A, R x (f0 x) /\ (forall x' : A, R x x' -> f0 x = f0 x')
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P
c:Prop -> bool
Hc:forall x : Prop, c x = true <-> x
class_of:=fun a y : A => c (R a y):A -> A -> bool
isclass:=fun f0 : A -> bool => exists x : A, f0 x = true:(A -> bool) -> Prop
class:={f0 : A -> bool | isclass f0}:Type
contains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Prop
f:class -> A
Hf:forall x : class, contains x (f x)
h:(A -> bool) -> A -> bool
Hh:forall f0 : A -> bool, (forall x : A, f0 x = h f0 x) /\ (forall g : A -> bool, (forall x : A, f0 x = g x) -> h f0 = h g)
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class

forall x : A, R x (f (f' x)) /\ (forall x' : A, R x x' -> f (f' x) = f (f' x'))
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P
c:Prop -> bool
Hc:forall x0 : Prop, c x0 = true <-> x0
class_of:=fun a y : A => c (R a y):A -> A -> bool
isclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Prop
class:={f0 : A -> bool | isclass f0}:Type
contains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Prop
f:class -> A
Hf:forall x0 : class, contains x0 (f x0)
h:(A -> bool) -> A -> bool
Hh:forall f0 : A -> bool, (forall x0 : A, f0 x0 = h f0 x0) /\ (forall g : A -> bool, (forall x0 : A, f0 x0 = g x0) -> h f0 = h g)
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class
x:A

R x (f (f' x)) /\ (forall x' : A, R x x' -> f (f' x) = f (f' x'))
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P
c:Prop -> bool
Hc:forall x0 : Prop, c x0 = true <-> x0
class_of:=fun a y : A => c (R a y):A -> A -> bool
isclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Prop
class:={f0 : A -> bool | isclass f0}:Type
contains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Prop
f:class -> A
Hf:forall x0 : class, contains x0 (f x0)
h:(A -> bool) -> A -> bool
Hh:forall f0 : A -> bool, (forall x0 : A, f0 x0 = h f0 x0) /\ (forall g : A -> bool, (forall x0 : A, f0 x0 = g x0) -> h f0 = h g)
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class
x:A
Hx:forall x0 : A, class_of x x0 = h (class_of x) x0
Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h g

R x (f (f' x)) /\ (forall x' : A, R x x' -> f (f' x) = f (f' x'))
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P
c:Prop -> bool
Hc:forall x0 : Prop, c x0 = true <-> x0
class_of:=fun a y : A => c (R a y):A -> A -> bool
isclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Prop
class:={f0 : A -> bool | isclass f0}:Type
contains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Prop
f:class -> A
Hf:forall x0 : class, contains x0 (f x0)
h:(A -> bool) -> A -> bool
Hh:forall f0 : A -> bool, (forall x0 : A, f0 x0 = h f0 x0) /\ (forall g : A -> bool, (forall x0 : A, f0 x0 = g x0) -> h f0 = h g)
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class
x:A
Hx:forall x0 : A, class_of x x0 = h (class_of x) x0
Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h g

R x (f (f' x))
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P
c:Prop -> bool
Hc:forall x0 : Prop, c x0 = true <-> x0
class_of:=fun a y : A => c (R a y):A -> A -> bool
isclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Prop
class:={f0 : A -> bool | isclass f0}:Type
contains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Prop
f:class -> A
Hf:forall x0 : class, contains x0 (f x0)
h:(A -> bool) -> A -> bool
Hh:forall f0 : A -> bool, (forall x0 : A, f0 x0 = h f0 x0) /\ (forall g : A -> bool, (forall x0 : A, f0 x0 = g x0) -> h f0 = h g)
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class
x:A
Hx:forall x0 : A, class_of x x0 = h (class_of x) x0
Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h g
forall x' : A, R x x' -> f (f' x) = f (f' x')
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P
c:Prop -> bool
Hc:forall x0 : Prop, c x0 = true <-> x0
class_of:=fun a y : A => c (R a y):A -> A -> bool
isclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Prop
class:={f0 : A -> bool | isclass f0}:Type
contains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Prop
f:class -> A
Hf:forall x0 : class, contains x0 (f x0)
h:(A -> bool) -> A -> bool
Hh:forall f0 : A -> bool, (forall x0 : A, f0 x0 = h f0 x0) /\ (forall g : A -> bool, (forall x0 : A, f0 x0 = g x0) -> h f0 = h g)
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class
x:A
Hx:forall x0 : A, class_of x x0 = h (class_of x) x0
Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h g

R x (f (f' x))
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P
c:Prop -> bool
Hc:forall x0 : Prop, c x0 = true <-> x0
class_of:=fun a y : A => c (R a y):A -> A -> bool
isclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Prop
class:={f0 : A -> bool | isclass f0}:Type
contains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Prop
f:class -> A
h:(A -> bool) -> A -> bool
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class
x:A
Hf:contains (f' x) (f (f' x))
Hh:forall f0 : A -> bool, (forall x0 : A, f0 x0 = h f0 x0) /\ (forall g : A -> bool, (forall x0 : A, f0 x0 = g x0) -> h f0 = h g)
Hx:forall x0 : A, class_of x x0 = h (class_of x) x0
Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h g

R x (f (f' x))
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P
c:Prop -> bool
Hc:forall x0 : Prop, c x0 = true <-> x0
class_of:=fun a y : A => c (R a y):A -> A -> bool
isclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Prop
class:={f0 : A -> bool | isclass f0}:Type
contains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Prop
f:class -> A
h:(A -> bool) -> A -> bool
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class
x:A
Hf:proj1_sig (f' x) (f (f' x)) = true
Hh:forall f0 : A -> bool, (forall x0 : A, f0 x0 = h f0 x0) /\ (forall g : A -> bool, (forall x0 : A, f0 x0 = g x0) -> h f0 = h g)
Hx:forall x0 : A, class_of x x0 = h (class_of x) x0
Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h g

R x (f (f' x))
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P
c:Prop -> bool
Hc:forall x0 : Prop, c x0 = true <-> x0
class_of:=fun a y : A => c (R a y):A -> A -> bool
isclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Prop
class:={f0 : A -> bool | isclass f0}:Type
contains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Prop
f:class -> A
h:(A -> bool) -> A -> bool
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class
x:A
Hf:h (class_of x) (f (f' x)) = true
Hh:forall f0 : A -> bool, (forall x0 : A, f0 x0 = h f0 x0) /\ (forall g : A -> bool, (forall x0 : A, f0 x0 = g x0) -> h f0 = h g)
Hx:forall x0 : A, class_of x x0 = h (class_of x) x0
Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h g

R x (f (f' x))
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P
c:Prop -> bool
Hc:forall x0 : Prop, c x0 = true <-> x0
class_of:=fun a y : A => c (R a y):A -> A -> bool
isclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Prop
class:={f0 : A -> bool | isclass f0}:Type
contains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Prop
f:class -> A
h:(A -> bool) -> A -> bool
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class
x:A
Hf:class_of x (f (f' x)) = true
Hh:forall f0 : A -> bool, (forall x0 : A, f0 x0 = h f0 x0) /\ (forall g : A -> bool, (forall x0 : A, f0 x0 = g x0) -> h f0 = h g)
Hx:forall x0 : A, class_of x x0 = h (class_of x) x0
Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h g

R x (f (f' x))
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P
c:Prop -> bool
Hc:forall x0 : Prop, c x0 = true <-> x0
class_of:=fun a y : A => c (R a y):A -> A -> bool
isclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Prop
class:={f0 : A -> bool | isclass f0}:Type
contains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Prop
f:class -> A
h:(A -> bool) -> A -> bool
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class
x:A
Hf:class_of x (f (f' x)) = true
Hh:forall f0 : A -> bool, (forall x0 : A, f0 x0 = h f0 x0) /\ (forall g : A -> bool, (forall x0 : A, f0 x0 = g x0) -> h f0 = h g)
Hx:forall x0 : A, class_of x x0 = h (class_of x) x0
Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h g

c (R x (f (f' x))) = true
assumption.
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P
c:Prop -> bool
Hc:forall x0 : Prop, c x0 = true <-> x0
class_of:=fun a y : A => c (R a y):A -> A -> bool
isclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Prop
class:={f0 : A -> bool | isclass f0}:Type
contains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Prop
f:class -> A
Hf:forall x0 : class, contains x0 (f x0)
h:(A -> bool) -> A -> bool
Hh:forall f0 : A -> bool, (forall x0 : A, f0 x0 = h f0 x0) /\ (forall g : A -> bool, (forall x0 : A, f0 x0 = g x0) -> h f0 = h g)
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class
x:A
Hx:forall x0 : A, class_of x x0 = h (class_of x) x0
Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h g

forall x' : A, R x x' -> f (f' x) = f (f' x')
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P
c:Prop -> bool
Hc:forall x0 : Prop, c x0 = true <-> x0
class_of:=fun a y0 : A => c (R a y0):A -> A -> bool
isclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Prop
class:={f0 : A -> bool | isclass f0}:Type
contains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Prop
f:class -> A
Hf:forall x0 : class, contains x0 (f x0)
h:(A -> bool) -> A -> bool
Hh:forall f0 : A -> bool, (forall x0 : A, f0 x0 = h f0 x0) /\ (forall g : A -> bool, (forall x0 : A, f0 x0 = g x0) -> h f0 = h g)
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class
x:A
Hx:forall x0 : A, class_of x x0 = h (class_of x) x0
Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h g
y:A
Hxy:R x y

f (f' x) = f (f' y)
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P
c:Prop -> bool
Hc:forall x0 : Prop, c x0 = true <-> x0
class_of:=fun a y0 : A => c (R a y0):A -> A -> bool
isclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Prop
class:={f0 : A -> bool | isclass f0}:Type
contains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Prop
f:class -> A
Hf:forall x0 : class, contains x0 (f x0)
h:(A -> bool) -> A -> bool
Hh:forall f0 : A -> bool, (forall x0 : A, f0 x0 = h f0 x0) /\ (forall g : A -> bool, (forall x0 : A, f0 x0 = g x0) -> h f0 = h g)
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class
x:A
Hx:forall x0 : A, class_of x x0 = h (class_of x) x0
Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h g
y:A
Hxy:R x y

f' x = f' y
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P
c:Prop -> bool
Hc:forall x0 : Prop, c x0 = true <-> x0
class_of:=fun a y0 : A => c (R a y0):A -> A -> bool
isclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Prop
class:={f0 : A -> bool | isclass f0}:Type
contains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Prop
f:class -> A
Hf:forall x0 : class, contains x0 (f x0)
h:(A -> bool) -> A -> bool
Hh:forall f0 : A -> bool, (forall x0 : A, f0 x0 = h f0 x0) /\ (forall g : A -> bool, (forall x0 : A, f0 x0 = g x0) -> h f0 = h g)
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class
x:A
Hx:forall x0 : A, class_of x x0 = h (class_of x) x0
Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h g
y:A
Hxy:R x y

h (class_of x) = h (class_of y)
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P
c:Prop -> bool
Hc:forall x0 : Prop, c x0 = true <-> x0
class_of:=fun a y0 : A => c (R a y0):A -> A -> bool
isclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Prop
class:={f0 : A -> bool | isclass f0}:Type
contains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Prop
f:class -> A
Hf:forall x0 : class, contains x0 (f x0)
h:(A -> bool) -> A -> bool
Hh:forall f0 : A -> bool, (forall x0 : A, f0 x0 = h f0 x0) /\ (forall g : A -> bool, (forall x0 : A, f0 x0 = g x0) -> h f0 = h g)
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class
x:A
Hx:forall x0 : A, class_of x x0 = h (class_of x) x0
Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h g
y:A
Hxy:R x y
Heq1:h (class_of x) = h (class_of y)
f' x = f' y
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P
c:Prop -> bool
Hc:forall x0 : Prop, c x0 = true <-> x0
class_of:=fun a y0 : A => c (R a y0):A -> A -> bool
isclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Prop
class:={f0 : A -> bool | isclass f0}:Type
contains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Prop
f:class -> A
Hf:forall x0 : class, contains x0 (f x0)
h:(A -> bool) -> A -> bool
Hh:forall f0 : A -> bool, (forall x0 : A, f0 x0 = h f0 x0) /\ (forall g : A -> bool, (forall x0 : A, f0 x0 = g x0) -> h f0 = h g)
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class
x:A
Hx:forall x0 : A, class_of x x0 = h (class_of x) x0
Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h g
y:A
Hxy:R x y

h (class_of x) = h (class_of y)
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P
c:Prop -> bool
Hc:forall x0 : Prop, c x0 = true <-> x0
class_of:=fun a y0 : A => c (R a y0):A -> A -> bool
isclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Prop
class:={f0 : A -> bool | isclass f0}:Type
contains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Prop
f:class -> A
Hf:forall x0 : class, contains x0 (f x0)
h:(A -> bool) -> A -> bool
Hh:forall f0 : A -> bool, (forall x0 : A, f0 x0 = h f0 x0) /\ (forall g : A -> bool, (forall x0 : A, f0 x0 = g x0) -> h f0 = h g)
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class
x:A
Hx:forall x0 : A, class_of x x0 = h (class_of x) x0
Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h g
y:A
Hxy:R x y

forall x0 : A, class_of x x0 = class_of y x0
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P
c:Prop -> bool
Hc:forall x0 : Prop, c x0 = true <-> x0
class_of:=fun a y0 : A => c (R a y0):A -> A -> bool
isclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Prop
class:={f0 : A -> bool | isclass f0}:Type
contains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Prop
f:class -> A
Hf:forall x0 : class, contains x0 (f x0)
h:(A -> bool) -> A -> bool
Hh:forall f0 : A -> bool, (forall x0 : A, f0 x0 = h f0 x0) /\ (forall g : A -> bool, (forall x0 : A, f0 x0 = g x0) -> h f0 = h g)
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class
x:A
Hx:forall x0 : A, class_of x x0 = h (class_of x) x0
Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h g
y:A
Hxy:R x y
z:A

class_of x z = class_of y z
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P
c:Prop -> bool
Hc:forall x0 : Prop, c x0 = true <-> x0
class_of:=fun a y0 : A => c (R a y0):A -> A -> bool
isclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Prop
class:={f0 : A -> bool | isclass f0}:Type
contains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Prop
f:class -> A
Hf:forall x0 : class, contains x0 (f x0)
h:(A -> bool) -> A -> bool
Hh:forall f0 : A -> bool, (forall x0 : A, f0 x0 = h f0 x0) /\ (forall g : A -> bool, (forall x0 : A, f0 x0 = g x0) -> h f0 = h g)
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class
x:A
Hx:forall x0 : A, class_of x x0 = h (class_of x) x0
Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h g
y:A
Hxy:R x y
z:A

c (R x z) = c (R y z)
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P
c:Prop -> bool
Hc:forall x0 : Prop, c x0 = true <-> x0
class_of:=fun a y0 : A => c (R a y0):A -> A -> bool
isclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Prop
class:={f0 : A -> bool | isclass f0}:Type
contains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Prop
f:class -> A
Hf:forall x0 : class, contains x0 (f x0)
h:(A -> bool) -> A -> bool
Hh:forall f0 : A -> bool, (forall x0 : A, f0 x0 = h f0 x0) /\ (forall g : A -> bool, (forall x0 : A, f0 x0 = g x0) -> h f0 = h g)
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class
x:A
Hx:forall x0 : A, class_of x x0 = h (class_of x) x0
Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h g
y:A
Hxy:R x y
z:A
Hxz:c (R x z) = true

true = c (R y z)
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P
c:Prop -> bool
Hc:forall x0 : Prop, c x0 = true <-> x0
class_of:=fun a y0 : A => c (R a y0):A -> A -> bool
isclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Prop
class:={f0 : A -> bool | isclass f0}:Type
contains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Prop
f:class -> A
Hf:forall x0 : class, contains x0 (f x0)
h:(A -> bool) -> A -> bool
Hh:forall f0 : A -> bool, (forall x0 : A, f0 x0 = h f0 x0) /\ (forall g : A -> bool, (forall x0 : A, f0 x0 = g x0) -> h f0 = h g)
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class
x:A
Hx:forall x0 : A, class_of x x0 = h (class_of x) x0
Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h g
y:A
Hxy:R x y
z:A
Hxz:c (R x z) = false
false = c (R y z)
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P
c:Prop -> bool
Hc:forall x0 : Prop, c x0 = true <-> x0
class_of:=fun a y0 : A => c (R a y0):A -> A -> bool
isclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Prop
class:={f0 : A -> bool | isclass f0}:Type
contains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Prop
f:class -> A
Hf:forall x0 : class, contains x0 (f x0)
h:(A -> bool) -> A -> bool
Hh:forall f0 : A -> bool, (forall x0 : A, f0 x0 = h f0 x0) /\ (forall g : A -> bool, (forall x0 : A, f0 x0 = g x0) -> h f0 = h g)
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class
x:A
Hx:forall x0 : A, class_of x x0 = h (class_of x) x0
Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h g
y:A
Hxy:R x y
z:A
Hxz:c (R x z) = true

true = c (R y z)
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P
c:Prop -> bool
Hc:forall x0 : Prop, c x0 = true <-> x0
class_of:=fun a y0 : A => c (R a y0):A -> A -> bool
isclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Prop
class:={f0 : A -> bool | isclass f0}:Type
contains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Prop
f:class -> A
Hf:forall x0 : class, contains x0 (f x0)
h:(A -> bool) -> A -> bool
Hh:forall f0 : A -> bool, (forall x0 : A, f0 x0 = h f0 x0) /\ (forall g : A -> bool, (forall x0 : A, f0 x0 = g x0) -> h f0 = h g)
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class
x:A
Hx:forall x0 : A, class_of x x0 = h (class_of x) x0
Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h g
y:A
Hxy:R x y
z:A
Hxz:c (R x z) = true

c (R y z) = true
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P
c:Prop -> bool
Hc:forall x0 : Prop, c x0 = true <-> x0
class_of:=fun a y0 : A => c (R a y0):A -> A -> bool
isclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Prop
class:={f0 : A -> bool | isclass f0}:Type
contains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Prop
f:class -> A
Hf:forall x0 : class, contains x0 (f x0)
h:(A -> bool) -> A -> bool
Hh:forall f0 : A -> bool, (forall x0 : A, f0 x0 = h f0 x0) /\ (forall g : A -> bool, (forall x0 : A, f0 x0 = g x0) -> h f0 = h g)
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class
x:A
Hx:forall x0 : A, class_of x x0 = h (class_of x) x0
Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h g
y:A
Hxy:R x y
z:A
Hxz:c (R x z) = true

R y z
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P
c:Prop -> bool
Hc:forall x0 : Prop, c x0 = true <-> x0
class_of:=fun a y0 : A => c (R a y0):A -> A -> bool
isclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Prop
class:={f0 : A -> bool | isclass f0}:Type
contains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Prop
f:class -> A
Hf:forall x0 : class, contains x0 (f x0)
h:(A -> bool) -> A -> bool
Hh:forall f0 : A -> bool, (forall x0 : A, f0 x0 = h f0 x0) /\ (forall g : A -> bool, (forall x0 : A, f0 x0 = g x0) -> h f0 = h g)
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class
x:A
Hx:forall x0 : A, class_of x x0 = h (class_of x) x0
Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h g
y:A
Hxy:R x y
z:A
Hxz:R x z

R y z
firstorder.
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P
c:Prop -> bool
Hc:forall x0 : Prop, c x0 = true <-> x0
class_of:=fun a y0 : A => c (R a y0):A -> A -> bool
isclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Prop
class:={f0 : A -> bool | isclass f0}:Type
contains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Prop
f:class -> A
Hf:forall x0 : class, contains x0 (f x0)
h:(A -> bool) -> A -> bool
Hh:forall f0 : A -> bool, (forall x0 : A, f0 x0 = h f0 x0) /\ (forall g : A -> bool, (forall x0 : A, f0 x0 = g x0) -> h f0 = h g)
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class
x:A
Hx:forall x0 : A, class_of x x0 = h (class_of x) x0
Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h g
y:A
Hxy:R x y
z:A
Hxz:c (R x z) = false

false = c (R y z)
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P
c:Prop -> bool
Hc:forall x0 : Prop, c x0 = true <-> x0
class_of:=fun a y0 : A => c (R a y0):A -> A -> bool
isclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Prop
class:={f0 : A -> bool | isclass f0}:Type
contains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Prop
f:class -> A
Hf:forall x0 : class, contains x0 (f x0)
h:(A -> bool) -> A -> bool
Hh:forall f0 : A -> bool, (forall x0 : A, f0 x0 = h f0 x0) /\ (forall g : A -> bool, (forall x0 : A, f0 x0 = g x0) -> h f0 = h g)
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class
x:A
Hx:forall x0 : A, class_of x x0 = h (class_of x) x0
Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h g
y:A
Hxy:R x y
z:A
Hxz:c (R x z) = false
Hyz:c (R y z) = true

false = true
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P
c:Prop -> bool
Hc:forall x0 : Prop, c x0 = true <-> x0
class_of:=fun a y0 : A => c (R a y0):A -> A -> bool
isclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Prop
class:={f0 : A -> bool | isclass f0}:Type
contains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Prop
f:class -> A
Hf:forall x0 : class, contains x0 (f x0)
h:(A -> bool) -> A -> bool
Hh:forall f0 : A -> bool, (forall x0 : A, f0 x0 = h f0 x0) /\ (forall g : A -> bool, (forall x0 : A, f0 x0 = g x0) -> h f0 = h g)
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class
x:A
Hx:forall x0 : A, class_of x x0 = h (class_of x) x0
Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h g
y:A
Hxy:R x y
z:A
Hxz:c (R x z) = false
Hyz:c (R y z) = false
false = false
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P
c:Prop -> bool
Hc:forall x0 : Prop, c x0 = true <-> x0
class_of:=fun a y0 : A => c (R a y0):A -> A -> bool
isclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Prop
class:={f0 : A -> bool | isclass f0}:Type
contains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Prop
f:class -> A
Hf:forall x0 : class, contains x0 (f x0)
h:(A -> bool) -> A -> bool
Hh:forall f0 : A -> bool, (forall x0 : A, f0 x0 = h f0 x0) /\ (forall g : A -> bool, (forall x0 : A, f0 x0 = g x0) -> h f0 = h g)
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class
x:A
Hx:forall x0 : A, class_of x x0 = h (class_of x) x0
Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h g
y:A
Hxy:R x y
z:A
Hxz:c (R x z) = false
Hyz:c (R y z) = true

false = true
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P
c:Prop -> bool
Hc:forall x0 : Prop, c x0 = true <-> x0
class_of:=fun a y0 : A => c (R a y0):A -> A -> bool
isclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Prop
class:={f0 : A -> bool | isclass f0}:Type
contains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Prop
f:class -> A
Hf:forall x0 : class, contains x0 (f x0)
h:(A -> bool) -> A -> bool
Hh:forall f0 : A -> bool, (forall x0 : A, f0 x0 = h f0 x0) /\ (forall g : A -> bool, (forall x0 : A, f0 x0 = g x0) -> h f0 = h g)
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class
x:A
Hx:forall x0 : A, class_of x x0 = h (class_of x) x0
Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h g
y:A
Hxy:R x y
z:A
Hxz:c (R x z) = false
Hyz:R y z

false = true
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P
c:Prop -> bool
Hc:forall x0 : Prop, c x0 = true <-> x0
class_of:=fun a y0 : A => c (R a y0):A -> A -> bool
isclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Prop
class:={f0 : A -> bool | isclass f0}:Type
contains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Prop
f:class -> A
Hf:forall x0 : class, contains x0 (f x0)
h:(A -> bool) -> A -> bool
Hh:forall f0 : A -> bool, (forall x0 : A, f0 x0 = h f0 x0) /\ (forall g : A -> bool, (forall x0 : A, f0 x0 = g x0) -> h f0 = h g)
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class
x:A
Hx:forall x0 : A, class_of x x0 = h (class_of x) x0
Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h g
y:A
Hxy:R x y
z:A
Hxz:c (R x z) = false
Hyz:R y z

c (R x z) = true
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P
c:Prop -> bool
Hc:forall x0 : Prop, c x0 = true <-> x0
class_of:=fun a y0 : A => c (R a y0):A -> A -> bool
isclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Prop
class:={f0 : A -> bool | isclass f0}:Type
contains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Prop
f:class -> A
Hf:forall x0 : class, contains x0 (f x0)
h:(A -> bool) -> A -> bool
Hh:forall f0 : A -> bool, (forall x0 : A, f0 x0 = h f0 x0) /\ (forall g : A -> bool, (forall x0 : A, f0 x0 = g x0) -> h f0 = h g)
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class
x:A
Hx:forall x0 : A, class_of x x0 = h (class_of x) x0
Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h g
y:A
Hxy:R x y
z:A
Hxz:c (R x z) = false
Hyz:R y z

R x z
firstorder.
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P
c:Prop -> bool
Hc:forall x0 : Prop, c x0 = true <-> x0
class_of:=fun a y0 : A => c (R a y0):A -> A -> bool
isclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Prop
class:={f0 : A -> bool | isclass f0}:Type
contains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Prop
f:class -> A
Hf:forall x0 : class, contains x0 (f x0)
h:(A -> bool) -> A -> bool
Hh:forall f0 : A -> bool, (forall x0 : A, f0 x0 = h f0 x0) /\ (forall g : A -> bool, (forall x0 : A, f0 x0 = g x0) -> h f0 = h g)
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class
x:A
Hx:forall x0 : A, class_of x x0 = h (class_of x) x0
Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h g
y:A
Hxy:R x y
z:A
Hxz:c (R x z) = false
Hyz:c (R y z) = false

false = false
easy.
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P
c:Prop -> bool
Hc:forall x0 : Prop, c x0 = true <-> x0
class_of:=fun a y0 : A => c (R a y0):A -> A -> bool
isclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Prop
class:={f0 : A -> bool | isclass f0}:Type
contains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Prop
f:class -> A
Hf:forall x0 : class, contains x0 (f x0)
h:(A -> bool) -> A -> bool
Hh:forall f0 : A -> bool, (forall x0 : A, f0 x0 = h f0 x0) /\ (forall g : A -> bool, (forall x0 : A, f0 x0 = g x0) -> h f0 = h g)
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class
x:A
Hx:forall x0 : A, class_of x x0 = h (class_of x) x0
Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h g
y:A
Hxy:R x y
Heq1:h (class_of x) = h (class_of y)

f' x = f' y
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P
c:Prop -> bool
Hc:forall x0 : Prop, c x0 = true <-> x0
class_of:=fun a y0 : A => c (R a y0):A -> A -> bool
isclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Prop
class:={f0 : A -> bool | isclass f0}:Type
contains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Prop
f:class -> A
Hf:forall x0 : class, contains x0 (f x0)
h:(A -> bool) -> A -> bool
Hh:forall f0 : A -> bool, (forall x0 : A, f0 x0 = h f0 x0) /\ (forall g : A -> bool, (forall x0 : A, f0 x0 = g x0) -> h f0 = h g)
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class
x:A
Hx:forall x0 : A, class_of x x0 = h (class_of x) x0
Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h g
y:A
Hxy:R x y
Heq1:h (class_of x) = h (class_of y)

rew [isclass] Heq1 in Hisclass x = Hisclass y
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P
c:Prop -> bool
Hc:forall x0 : Prop, c x0 = true <-> x0
class_of:=fun a y0 : A => c (R a y0):A -> A -> bool
isclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Prop
class:={f0 : A -> bool | isclass f0}:Type
contains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Prop
f:class -> A
Hf:forall x0 : class, contains x0 (f x0)
h:(A -> bool) -> A -> bool
Hh:forall f0 : A -> bool, (forall x0 : A, f0 x0 = h f0 x0) /\ (forall g : A -> bool, (forall x0 : A, f0 x0 = g x0) -> h f0 = h g)
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class
x:A
Hx:forall x0 : A, class_of x x0 = h (class_of x) x0
Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h g
y:A
Hxy:R x y
Heq1:h (class_of x) = h (class_of y)
Heq2:rew [isclass] Heq1 in Hisclass x = Hisclass y
f' x = f' y
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P
c:Prop -> bool
Hc:forall x0 : Prop, c x0 = true <-> x0
class_of:=fun a y0 : A => c (R a y0):A -> A -> bool
isclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Prop
class:={f0 : A -> bool | isclass f0}:Type
contains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Prop
f:class -> A
Hf:forall x0 : class, contains x0 (f x0)
h:(A -> bool) -> A -> bool
Hh:forall f0 : A -> bool, (forall x0 : A, f0 x0 = h f0 x0) /\ (forall g : A -> bool, (forall x0 : A, f0 x0 = g x0) -> h f0 = h g)
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class
x:A
Hx:forall x0 : A, class_of x x0 = h (class_of x) x0
Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h g
y:A
Hxy:R x y
Heq1:h (class_of x) = h (class_of y)

rew [isclass] Heq1 in Hisclass x = Hisclass y
apply proof_irrelevance_cci, EM.
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P
c:Prop -> bool
Hc:forall x0 : Prop, c x0 = true <-> x0
class_of:=fun a y0 : A => c (R a y0):A -> A -> bool
isclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Prop
class:={f0 : A -> bool | isclass f0}:Type
contains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Prop
f:class -> A
Hf:forall x0 : class, contains x0 (f x0)
h:(A -> bool) -> A -> bool
Hh:forall f0 : A -> bool, (forall x0 : A, f0 x0 = h f0 x0) /\ (forall g : A -> bool, (forall x0 : A, f0 x0 = g x0) -> h f0 = h g)
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class
x:A
Hx:forall x0 : A, class_of x x0 = h (class_of x) x0
Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h g
y:A
Hxy:R x y
Heq1:h (class_of x) = h (class_of y)
Heq2:rew [isclass] Heq1 in Hisclass x = Hisclass y

f' x = f' y
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P
c:Prop -> bool
Hc:forall x0 : Prop, c x0 = true <-> x0
class_of:=fun a y0 : A => c (R a y0):A -> A -> bool
isclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Prop
class:={f0 : A -> bool | isclass f0}:Type
contains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Prop
f:class -> A
Hf:forall x0 : class, contains x0 (f x0)
h:(A -> bool) -> A -> bool
Hh:forall f0 : A -> bool, (forall x0 : A, f0 x0 = h f0 x0) /\ (forall g : A -> bool, (forall x0 : A, f0 x0 = g x0) -> h f0 = h g)
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class
x:A
Hx:forall x0 : A, class_of x x0 = h (class_of x) x0
Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h g
y:A
Hxy:R x y
Heq1:h (class_of x) = h (class_of y)
Heq2:rew [isclass] Heq1 in Hisclass x = Hisclass y

exist isclass (h (class_of x)) (Hisclass x) = exist isclass (h (class_of y)) (Hisclass y)
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P
c:Prop -> bool
Hc:forall x0 : Prop, c x0 = true <-> x0
class_of:=fun a y0 : A => c (R a y0):A -> A -> bool
isclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Prop
class:={f0 : A -> bool | isclass f0}:Type
contains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Prop
f:class -> A
Hf:forall x0 : class, contains x0 (f x0)
h:(A -> bool) -> A -> bool
Hh:forall f0 : A -> bool, (forall x0 : A, f0 x0 = h f0 x0) /\ (forall g : A -> bool, (forall x0 : A, f0 x0 = g x0) -> h f0 = h g)
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class
x:A
Hx:forall x0 : A, class_of x x0 = h (class_of x) x0
Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h g
y:A
Hxy:R x y
Heq1:h (class_of x) = h (class_of y)
Heq2:rew [isclass] Heq1 in Hisclass x = Hisclass y

exist isclass (h (class_of x)) (Hisclass x) = exist isclass (h (class_of y)) (rew [isclass] Heq1 in Hisclass x)
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
H:forall P : Prop, exists b : bool, b = true <-> P
c:Prop -> bool
Hc:forall x0 : Prop, c x0 = true <-> x0
class_of:=fun a y0 : A => c (R a y0):A -> A -> bool
isclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Prop
class:={f0 : A -> bool | isclass f0}:Type
contains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Prop
f:class -> A
Hf:forall x0 : class, contains x0 (f x0)
h:(A -> bool) -> A -> bool
Hh:forall f0 : A -> bool, (forall x0 : A, f0 x0 = h f0 x0) /\ (forall g : A -> bool, (forall x0 : A, f0 x0 = g x0) -> h f0 = h g)
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class
x:A
Hx:forall x0 : A, class_of x x0 = h (class_of x) x0
Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h g
y:A
Hxy:R x y
Heq1:h (class_of x) = h (class_of y)
Heq2:rew [isclass] Heq1 in Hisclass x = Hisclass y

exist isclass (h (class_of x)) (Hisclass x) = exist isclass (h (class_of x)) (rew [isclass] eq_refl in Hisclass x)
reflexivity. Qed.

FunctionalChoice /\ ExtensionalFunctionRepresentative /\ ExcludedMiddle <-> SetoidFunctionalChoice

FunctionalChoice /\ ExtensionalFunctionRepresentative /\ ExcludedMiddle <-> SetoidFunctionalChoice

FunctionalChoice /\ ExtensionalFunctionRepresentative /\ ExcludedMiddle -> SetoidFunctionalChoice

SetoidFunctionalChoice -> FunctionalChoice /\ ExtensionalFunctionRepresentative /\ ExcludedMiddle

FunctionalChoice /\ ExtensionalFunctionRepresentative /\ ExcludedMiddle -> SetoidFunctionalChoice
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle

SetoidFunctionalChoice
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle

FunctionalRelReification
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
RepresentativeFunctionalChoice
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle

FunctionalRelReification
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle
A, B:Type

FunctionalRelReification_on A B
apply fun_choice_imp_functional_rel_reification, FunChoice.
FunChoice:FunctionalChoice
SetoidFunRepr:ExtensionalFunctionRepresentative
EM:ExcludedMiddle

RepresentativeFunctionalChoice
now apply fun_choice_and_ext_functions_repr_and_excluded_middle_imp_setoid_fun_choice.

SetoidFunctionalChoice -> FunctionalChoice /\ ExtensionalFunctionRepresentative /\ ExcludedMiddle
SetoidFunChoice:SetoidFunctionalChoice

FunctionalChoice /\ ExtensionalFunctionRepresentative /\ ExcludedMiddle
SetoidFunChoice:SetoidFunctionalChoice

FunctionalChoice
SetoidFunChoice:SetoidFunctionalChoice
ExtensionalFunctionRepresentative
SetoidFunChoice:SetoidFunctionalChoice
ExcludedMiddle
SetoidFunChoice:SetoidFunctionalChoice

FunctionalChoice
now intros A B; apply setoid_fun_choice_imp_fun_choice.
SetoidFunChoice:SetoidFunctionalChoice

ExtensionalFunctionRepresentative
SetoidFunChoice:SetoidFunctionalChoice

RepresentativeFunctionalChoice
now apply setoid_fun_choice_imp_repr_fun_choice.
SetoidFunChoice:SetoidFunctionalChoice

ExcludedMiddle
SetoidFunChoice:SetoidFunctionalChoice

RepresentativeFunctionalChoice
now apply setoid_fun_choice_imp_repr_fun_choice. Qed. (**********************************************************************)

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

FunctionalChoice -> ExtensionalPredicateRepresentative -> ProofIrrelevance -> RepresentativeFunctionalChoice

FunctionalChoice -> ExtensionalPredicateRepresentative -> ProofIrrelevance -> RepresentativeFunctionalChoice
FunChoice:FunctionalChoice
PredExtRepr:ExtensionalPredicateRepresentative
PI:ProofIrrelevance
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R

exists f : A -> A, forall x : A, R x (f x) /\ (forall x' : A, R x x' -> f x = f x')
FunChoice:FunctionalChoice
PredExtRepr:ExtensionalPredicateRepresentative
PI:ProofIrrelevance
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
isclass:=fun P : A -> Prop => exists x : A, P x:(A -> Prop) -> Prop

exists f : A -> A, forall x : A, R x (f x) /\ (forall x' : A, R x x' -> f x = f x')
FunChoice:FunctionalChoice
PredExtRepr:ExtensionalPredicateRepresentative
PI:ProofIrrelevance
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
isclass:=fun P : A -> Prop => exists x : A, P x:(A -> Prop) -> Prop
class:={P : A -> Prop | isclass P}:Type

exists f : A -> A, forall x : A, R x (f x) /\ (forall x' : A, R x x' -> f x = f x')
FunChoice:FunctionalChoice
PredExtRepr:ExtensionalPredicateRepresentative
PI:ProofIrrelevance
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
isclass:=fun P : A -> Prop => exists x : A, P x:(A -> Prop) -> Prop
class:={P : A -> Prop | isclass P}:Type
contains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Prop

exists f : A -> A, forall x : A, R x (f x) /\ (forall x' : A, R x x' -> f x = f x')
FunChoice:FunctionalChoice
PredExtRepr:ExtensionalPredicateRepresentative
PI:ProofIrrelevance
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
isclass:=fun P : A -> Prop => exists x : A, P x:(A -> Prop) -> Prop
class:={P : A -> Prop | isclass P}:Type
contains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Prop
class_of:=fun a : A => R a:A -> A -> Prop

exists f : A -> A, forall x : A, R x (f x) /\ (forall x' : A, R x x' -> f x = f x')
FunChoice:FunctionalChoice
PredExtRepr:ExtensionalPredicateRepresentative
PI:ProofIrrelevance
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
isclass:=fun P : A -> Prop => exists x : A, P x:(A -> Prop) -> Prop
class:={P : A -> Prop | isclass P}:Type
contains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Prop
class_of:=fun a : A => R a:A -> A -> Prop

forall x : class, exists y : A, contains x y
FunChoice:FunctionalChoice
PredExtRepr:ExtensionalPredicateRepresentative
PI:ProofIrrelevance
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
isclass:=fun P : A -> Prop => exists x : A, P x:(A -> Prop) -> Prop
class:={P : A -> Prop | isclass P}:Type
contains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Prop
class_of:=fun a : A => R a:A -> A -> Prop
f:class -> A
Hf:forall x : class, contains x (f x)
exists f0 : A -> A, forall x : A, R x (f0 x) /\ (forall x' : A, R x x' -> f0 x = f0 x')
FunChoice:FunctionalChoice
PredExtRepr:ExtensionalPredicateRepresentative
PI:ProofIrrelevance
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
isclass:=fun P : A -> Prop => exists x : A, P x:(A -> Prop) -> Prop
class:={P : A -> Prop | isclass P}:Type
contains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Prop
class_of:=fun a : A => R a:A -> A -> Prop

forall x : class, exists y : A, contains x y
FunChoice:FunctionalChoice
PredExtRepr:ExtensionalPredicateRepresentative
PI:ProofIrrelevance
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
isclass:=fun P : A -> Prop => exists x : A, P x:(A -> Prop) -> Prop
class:={P : A -> Prop | isclass P}:Type
contains:=fun (c0 : class) (a : A) => proj1_sig c0 a:class -> A -> Prop
class_of:=fun a : A => R a:A -> A -> Prop
c:class

exists y : A, contains c y
apply proj2_sig.
FunChoice:FunctionalChoice
PredExtRepr:ExtensionalPredicateRepresentative
PI:ProofIrrelevance
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
isclass:=fun P : A -> Prop => exists x : A, P x:(A -> Prop) -> Prop
class:={P : A -> Prop | isclass P}:Type
contains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Prop
class_of:=fun a : A => R a:A -> A -> Prop
f:class -> A
Hf:forall x : class, contains x (f x)

exists f0 : A -> A, forall x : A, R x (f0 x) /\ (forall x' : A, R x x' -> f0 x = f0 x')
FunChoice:FunctionalChoice
PredExtRepr:ExtensionalPredicateRepresentative
PI:ProofIrrelevance
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
isclass:=fun P : A -> Prop => exists x : A, P x:(A -> Prop) -> Prop
class:={P : A -> Prop | isclass P}:Type
contains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Prop
class_of:=fun a : A => R a:A -> A -> Prop
f:class -> A
Hf:forall x : class, contains x (f x)
h:(A -> Prop) -> A -> Prop
Hh:forall P : A -> Prop, (forall x : A, P x <-> h P x) /\ (forall Q : A -> Prop, (forall x : A, P x <-> Q x) -> h P = h Q)

exists f0 : A -> A, forall x : A, R x (f0 x) /\ (forall x' : A, R x x' -> f0 x = f0 x')
FunChoice:FunctionalChoice
PredExtRepr:ExtensionalPredicateRepresentative
PI:ProofIrrelevance
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
isclass:=fun P : A -> Prop => exists x : A, P x:(A -> Prop) -> Prop
class:={P : A -> Prop | isclass P}:Type
contains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Prop
class_of:=fun a : A => R a:A -> A -> Prop
f:class -> A
Hf:forall x : class, contains x (f x)
h:(A -> Prop) -> A -> Prop
Hh:forall P : A -> Prop, (forall x : A, P x <-> h P x) /\ (forall Q : A -> Prop, (forall x : A, P x <-> Q x) -> h P = h Q)

forall a : A, isclass (h (class_of a))
FunChoice:FunctionalChoice
PredExtRepr:ExtensionalPredicateRepresentative
PI:ProofIrrelevance
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
isclass:=fun P : A -> Prop => exists x : A, P x:(A -> Prop) -> Prop
class:={P : A -> Prop | isclass P}:Type
contains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Prop
class_of:=fun a : A => R a:A -> A -> Prop
f:class -> A
Hf:forall x : class, contains x (f x)
h:(A -> Prop) -> A -> Prop
Hh:forall P : A -> Prop, (forall x : A, P x <-> h P x) /\ (forall Q : A -> Prop, (forall x : A, P x <-> Q x) -> h P = h Q)
Hisclass:forall a : A, isclass (h (class_of a))
exists f0 : A -> A, forall x : A, R x (f0 x) /\ (forall x' : A, R x x' -> f0 x = f0 x')
FunChoice:FunctionalChoice
PredExtRepr:ExtensionalPredicateRepresentative
PI:ProofIrrelevance
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
isclass:=fun P : A -> Prop => exists x : A, P x:(A -> Prop) -> Prop
class:={P : A -> Prop | isclass P}:Type
contains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Prop
class_of:=fun a : A => R a:A -> A -> Prop
f:class -> A
Hf:forall x : class, contains x (f x)
h:(A -> Prop) -> A -> Prop
Hh:forall P : A -> Prop, (forall x : A, P x <-> h P x) /\ (forall Q : A -> Prop, (forall x : A, P x <-> Q x) -> h P = h Q)

forall a : A, isclass (h (class_of a))
FunChoice:FunctionalChoice
PredExtRepr:ExtensionalPredicateRepresentative
PI:ProofIrrelevance
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
isclass:=fun P : A -> Prop => exists x : A, P x:(A -> Prop) -> Prop
class:={P : A -> Prop | isclass P}:Type
contains:=fun (c : class) (a0 : A) => proj1_sig c a0:class -> A -> Prop
class_of:=fun a0 : A => R a0:A -> A -> Prop
f:class -> A
Hf:forall x : class, contains x (f x)
h:(A -> Prop) -> A -> Prop
Hh:forall P : A -> Prop, (forall x : A, P x <-> h P x) /\ (forall Q : A -> Prop, (forall x : A, P x <-> Q x) -> h P = h Q)
a:A

isclass (h (class_of a))
FunChoice:FunctionalChoice
PredExtRepr:ExtensionalPredicateRepresentative
PI:ProofIrrelevance
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
isclass:=fun P : A -> Prop => exists x : A, P x:(A -> Prop) -> Prop
class:={P : A -> Prop | isclass P}:Type
contains:=fun (c : class) (a0 : A) => proj1_sig c a0:class -> A -> Prop
class_of:=fun a0 : A => R a0:A -> A -> Prop
f:class -> A
Hf:forall x : class, contains x (f x)
h:(A -> Prop) -> A -> Prop
Hh:forall P : A -> Prop, (forall x : A, P x <-> h P x) /\ (forall Q : A -> Prop, (forall x : A, P x <-> Q x) -> h P = h Q)
a:A

h (class_of a) a
FunChoice:FunctionalChoice
PredExtRepr:ExtensionalPredicateRepresentative
PI:ProofIrrelevance
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
isclass:=fun P : A -> Prop => exists x : A, P x:(A -> Prop) -> Prop
class:={P : A -> Prop | isclass P}:Type
contains:=fun (c : class) (a0 : A) => proj1_sig c a0:class -> A -> Prop
class_of:=fun a0 : A => R a0:A -> A -> Prop
f:class -> A
Hf:forall x : class, contains x (f x)
h:(A -> Prop) -> A -> Prop
Hh:forall P : A -> Prop, (forall x : A, P x <-> h P x) /\ (forall Q : A -> Prop, (forall x : A, P x <-> Q x) -> h P = h Q)
a:A
Ha:forall x : A, class_of a x <-> h (class_of a) x
Huniqa:forall Q : A -> Prop, (forall x : A, class_of a x <-> Q x) -> h (class_of a) = h Q

h (class_of a) a
rewrite <- Ha; apply Hrefl.
FunChoice:FunctionalChoice
PredExtRepr:ExtensionalPredicateRepresentative
PI:ProofIrrelevance
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
isclass:=fun P : A -> Prop => exists x : A, P x:(A -> Prop) -> Prop
class:={P : A -> Prop | isclass P}:Type
contains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Prop
class_of:=fun a : A => R a:A -> A -> Prop
f:class -> A
Hf:forall x : class, contains x (f x)
h:(A -> Prop) -> A -> Prop
Hh:forall P : A -> Prop, (forall x : A, P x <-> h P x) /\ (forall Q : A -> Prop, (forall x : A, P x <-> Q x) -> h P = h Q)
Hisclass:forall a : A, isclass (h (class_of a))

exists f0 : A -> A, forall x : A, R x (f0 x) /\ (forall x' : A, R x x' -> f0 x = f0 x')
FunChoice:FunctionalChoice
PredExtRepr:ExtensionalPredicateRepresentative
PI:ProofIrrelevance
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
isclass:=fun P : A -> Prop => exists x : A, P x:(A -> Prop) -> Prop
class:={P : A -> Prop | isclass P}:Type
contains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Prop
class_of:=fun a : A => R a:A -> A -> Prop
f:class -> A
Hf:forall x : class, contains x (f x)
h:(A -> Prop) -> A -> Prop
Hh:forall P : A -> Prop, (forall x : A, P x <-> h P x) /\ (forall Q : A -> Prop, (forall x : A, P x <-> Q x) -> h P = h Q)
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class

exists f0 : A -> A, forall x : A, R x (f0 x) /\ (forall x' : A, R x x' -> f0 x = f0 x')
FunChoice:FunctionalChoice
PredExtRepr:ExtensionalPredicateRepresentative
PI:ProofIrrelevance
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
isclass:=fun P : A -> Prop => exists x : A, P x:(A -> Prop) -> Prop
class:={P : A -> Prop | isclass P}:Type
contains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Prop
class_of:=fun a : A => R a:A -> A -> Prop
f:class -> A
Hf:forall x : class, contains x (f x)
h:(A -> Prop) -> A -> Prop
Hh:forall P : A -> Prop, (forall x : A, P x <-> h P x) /\ (forall Q : A -> Prop, (forall x : A, P x <-> Q x) -> h P = h Q)
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class

forall x : A, R x (f (f' x)) /\ (forall x' : A, R x x' -> f (f' x) = f (f' x'))
FunChoice:FunctionalChoice
PredExtRepr:ExtensionalPredicateRepresentative
PI:ProofIrrelevance
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
isclass:=fun P : A -> Prop => exists x0 : A, P x0:(A -> Prop) -> Prop
class:={P : A -> Prop | isclass P}:Type
contains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Prop
class_of:=fun a : A => R a:A -> A -> Prop
f:class -> A
Hf:forall x0 : class, contains x0 (f x0)
h:(A -> Prop) -> A -> Prop
Hh:forall P : A -> Prop, (forall x0 : A, P x0 <-> h P x0) /\ (forall Q : A -> Prop, (forall x0 : A, P x0 <-> Q x0) -> h P = h Q)
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class
x:A

R x (f (f' x)) /\ (forall x' : A, R x x' -> f (f' x) = f (f' x'))
FunChoice:FunctionalChoice
PredExtRepr:ExtensionalPredicateRepresentative
PI:ProofIrrelevance
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
isclass:=fun P : A -> Prop => exists x0 : A, P x0:(A -> Prop) -> Prop
class:={P : A -> Prop | isclass P}:Type
contains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Prop
class_of:=fun a : A => R a:A -> A -> Prop
f:class -> A
Hf:forall x0 : class, contains x0 (f x0)
h:(A -> Prop) -> A -> Prop
Hh:forall P : A -> Prop, (forall x0 : A, P x0 <-> h P x0) /\ (forall Q : A -> Prop, (forall x0 : A, P x0 <-> Q x0) -> h P = h Q)
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class
x:A
Hx:forall x0 : A, class_of x x0 <-> h (class_of x) x0
Huniqx:forall Q : A -> Prop, (forall x0 : A, class_of x x0 <-> Q x0) -> h (class_of x) = h Q

R x (f (f' x)) /\ (forall x' : A, R x x' -> f (f' x) = f (f' x'))
FunChoice:FunctionalChoice
PredExtRepr:ExtensionalPredicateRepresentative
PI:ProofIrrelevance
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
isclass:=fun P : A -> Prop => exists x0 : A, P x0:(A -> Prop) -> Prop
class:={P : A -> Prop | isclass P}:Type
contains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Prop
class_of:=fun a : A => R a:A -> A -> Prop
f:class -> A
Hf:forall x0 : class, contains x0 (f x0)
h:(A -> Prop) -> A -> Prop
Hh:forall P : A -> Prop, (forall x0 : A, P x0 <-> h P x0) /\ (forall Q : A -> Prop, (forall x0 : A, P x0 <-> Q x0) -> h P = h Q)
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class
x:A
Hx:forall x0 : A, class_of x x0 <-> h (class_of x) x0
Huniqx:forall Q : A -> Prop, (forall x0 : A, class_of x x0 <-> Q x0) -> h (class_of x) = h Q

R x (f (f' x))
FunChoice:FunctionalChoice
PredExtRepr:ExtensionalPredicateRepresentative
PI:ProofIrrelevance
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
isclass:=fun P : A -> Prop => exists x0 : A, P x0:(A -> Prop) -> Prop
class:={P : A -> Prop | isclass P}:Type
contains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Prop
class_of:=fun a : A => R a:A -> A -> Prop
f:class -> A
Hf:forall x0 : class, contains x0 (f x0)
h:(A -> Prop) -> A -> Prop
Hh:forall P : A -> Prop, (forall x0 : A, P x0 <-> h P x0) /\ (forall Q : A -> Prop, (forall x0 : A, P x0 <-> Q x0) -> h P = h Q)
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class
x:A
Hx:forall x0 : A, class_of x x0 <-> h (class_of x) x0
Huniqx:forall Q : A -> Prop, (forall x0 : A, class_of x x0 <-> Q x0) -> h (class_of x) = h Q
forall x' : A, R x x' -> f (f' x) = f (f' x')
FunChoice:FunctionalChoice
PredExtRepr:ExtensionalPredicateRepresentative
PI:ProofIrrelevance
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
isclass:=fun P : A -> Prop => exists x0 : A, P x0:(A -> Prop) -> Prop
class:={P : A -> Prop | isclass P}:Type
contains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Prop
class_of:=fun a : A => R a:A -> A -> Prop
f:class -> A
Hf:forall x0 : class, contains x0 (f x0)
h:(A -> Prop) -> A -> Prop
Hh:forall P : A -> Prop, (forall x0 : A, P x0 <-> h P x0) /\ (forall Q : A -> Prop, (forall x0 : A, P x0 <-> Q x0) -> h P = h Q)
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class
x:A
Hx:forall x0 : A, class_of x x0 <-> h (class_of x) x0
Huniqx:forall Q : A -> Prop, (forall x0 : A, class_of x x0 <-> Q x0) -> h (class_of x) = h Q

R x (f (f' x))
FunChoice:FunctionalChoice
PredExtRepr:ExtensionalPredicateRepresentative
PI:ProofIrrelevance
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
isclass:=fun P : A -> Prop => exists x0 : A, P x0:(A -> Prop) -> Prop
class:={P : A -> Prop | isclass P}:Type
contains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Prop
class_of:=fun a : A => R a:A -> A -> Prop
f:class -> A
h:(A -> Prop) -> A -> Prop
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class
x:A
Hf:contains (f' x) (f (f' x))
Hh:forall P : A -> Prop, (forall x0 : A, P x0 <-> h P x0) /\ (forall Q : A -> Prop, (forall x0 : A, P x0 <-> Q x0) -> h P = h Q)
Hx:forall x0 : A, class_of x x0 <-> h (class_of x) x0
Huniqx:forall Q : A -> Prop, (forall x0 : A, class_of x x0 <-> Q x0) -> h (class_of x) = h Q

R x (f (f' x))
FunChoice:FunctionalChoice
PredExtRepr:ExtensionalPredicateRepresentative
PI:ProofIrrelevance
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
isclass:=fun P : A -> Prop => exists x0 : A, P x0:(A -> Prop) -> Prop
class:={P : A -> Prop | isclass P}:Type
contains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Prop
class_of:=fun a : A => R a:A -> A -> Prop
f:class -> A
h:(A -> Prop) -> A -> Prop
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class
x:A
Hf:h (class_of x) (f (f' x))
Hh:forall P : A -> Prop, (forall x0 : A, P x0 <-> h P x0) /\ (forall Q : A -> Prop, (forall x0 : A, P x0 <-> Q x0) -> h P = h Q)
Hx:forall x0 : A, class_of x x0 <-> h (class_of x) x0
Huniqx:forall Q : A -> Prop, (forall x0 : A, class_of x x0 <-> Q x0) -> h (class_of x) = h Q

R x (f (f' x))
FunChoice:FunctionalChoice
PredExtRepr:ExtensionalPredicateRepresentative
PI:ProofIrrelevance
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
isclass:=fun P : A -> Prop => exists x0 : A, P x0:(A -> Prop) -> Prop
class:={P : A -> Prop | isclass P}:Type
contains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Prop
class_of:=fun a : A => R a:A -> A -> Prop
f:class -> A
h:(A -> Prop) -> A -> Prop
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class
x:A
Hf:class_of x (f (f' x))
Hh:forall P : A -> Prop, (forall x0 : A, P x0 <-> h P x0) /\ (forall Q : A -> Prop, (forall x0 : A, P x0 <-> Q x0) -> h P = h Q)
Hx:forall x0 : A, class_of x x0 <-> h (class_of x) x0
Huniqx:forall Q : A -> Prop, (forall x0 : A, class_of x x0 <-> Q x0) -> h (class_of x) = h Q

R x (f (f' x))
assumption.
FunChoice:FunctionalChoice
PredExtRepr:ExtensionalPredicateRepresentative
PI:ProofIrrelevance
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
isclass:=fun P : A -> Prop => exists x0 : A, P x0:(A -> Prop) -> Prop
class:={P : A -> Prop | isclass P}:Type
contains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Prop
class_of:=fun a : A => R a:A -> A -> Prop
f:class -> A
Hf:forall x0 : class, contains x0 (f x0)
h:(A -> Prop) -> A -> Prop
Hh:forall P : A -> Prop, (forall x0 : A, P x0 <-> h P x0) /\ (forall Q : A -> Prop, (forall x0 : A, P x0 <-> Q x0) -> h P = h Q)
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class
x:A
Hx:forall x0 : A, class_of x x0 <-> h (class_of x) x0
Huniqx:forall Q : A -> Prop, (forall x0 : A, class_of x x0 <-> Q x0) -> h (class_of x) = h Q

forall x' : A, R x x' -> f (f' x) = f (f' x')
FunChoice:FunctionalChoice
PredExtRepr:ExtensionalPredicateRepresentative
PI:ProofIrrelevance
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
isclass:=fun P : A -> Prop => exists x0 : A, P x0:(A -> Prop) -> Prop
class:={P : A -> Prop | isclass P}:Type
contains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Prop
class_of:=fun a : A => R a:A -> A -> Prop
f:class -> A
Hf:forall x0 : class, contains x0 (f x0)
h:(A -> Prop) -> A -> Prop
Hh:forall P : A -> Prop, (forall x0 : A, P x0 <-> h P x0) /\ (forall Q : A -> Prop, (forall x0 : A, P x0 <-> Q x0) -> h P = h Q)
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class
x:A
Hx:forall x0 : A, class_of x x0 <-> h (class_of x) x0
Huniqx:forall Q : A -> Prop, (forall x0 : A, class_of x x0 <-> Q x0) -> h (class_of x) = h Q
y:A
Hxy:R x y

f (f' x) = f (f' y)
FunChoice:FunctionalChoice
PredExtRepr:ExtensionalPredicateRepresentative
PI:ProofIrrelevance
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
isclass:=fun P : A -> Prop => exists x0 : A, P x0:(A -> Prop) -> Prop
class:={P : A -> Prop | isclass P}:Type
contains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Prop
class_of:=fun a : A => R a:A -> A -> Prop
f:class -> A
Hf:forall x0 : class, contains x0 (f x0)
h:(A -> Prop) -> A -> Prop
Hh:forall P : A -> Prop, (forall x0 : A, P x0 <-> h P x0) /\ (forall Q : A -> Prop, (forall x0 : A, P x0 <-> Q x0) -> h P = h Q)
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class
x:A
Hx:forall x0 : A, class_of x x0 <-> h (class_of x) x0
Huniqx:forall Q : A -> Prop, (forall x0 : A, class_of x x0 <-> Q x0) -> h (class_of x) = h Q
y:A
Hxy:R x y

f' x = f' y
FunChoice:FunctionalChoice
PredExtRepr:ExtensionalPredicateRepresentative
PI:ProofIrrelevance
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
isclass:=fun P : A -> Prop => exists x0 : A, P x0:(A -> Prop) -> Prop
class:={P : A -> Prop | isclass P}:Type
contains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Prop
class_of:=fun a : A => R a:A -> A -> Prop
f:class -> A
Hf:forall x0 : class, contains x0 (f x0)
h:(A -> Prop) -> A -> Prop
Hh:forall P : A -> Prop, (forall x0 : A, P x0 <-> h P x0) /\ (forall Q : A -> Prop, (forall x0 : A, P x0 <-> Q x0) -> h P = h Q)
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class
x:A
Hx:forall x0 : A, class_of x x0 <-> h (class_of x) x0
Huniqx:forall Q : A -> Prop, (forall x0 : A, class_of x x0 <-> Q x0) -> h (class_of x) = h Q
y:A
Hxy:R x y

h (class_of x) = h (class_of y)
FunChoice:FunctionalChoice
PredExtRepr:ExtensionalPredicateRepresentative
PI:ProofIrrelevance
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
isclass:=fun P : A -> Prop => exists x0 : A, P x0:(A -> Prop) -> Prop
class:={P : A -> Prop | isclass P}:Type
contains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Prop
class_of:=fun a : A => R a:A -> A -> Prop
f:class -> A
Hf:forall x0 : class, contains x0 (f x0)
h:(A -> Prop) -> A -> Prop
Hh:forall P : A -> Prop, (forall x0 : A, P x0 <-> h P x0) /\ (forall Q : A -> Prop, (forall x0 : A, P x0 <-> Q x0) -> h P = h Q)
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class
x:A
Hx:forall x0 : A, class_of x x0 <-> h (class_of x) x0
Huniqx:forall Q : A -> Prop, (forall x0 : A, class_of x x0 <-> Q x0) -> h (class_of x) = h Q
y:A
Hxy:R x y
Heq1:h (class_of x) = h (class_of y)
f' x = f' y
FunChoice:FunctionalChoice
PredExtRepr:ExtensionalPredicateRepresentative
PI:ProofIrrelevance
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
isclass:=fun P : A -> Prop => exists x0 : A, P x0:(A -> Prop) -> Prop
class:={P : A -> Prop | isclass P}:Type
contains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Prop
class_of:=fun a : A => R a:A -> A -> Prop
f:class -> A
Hf:forall x0 : class, contains x0 (f x0)
h:(A -> Prop) -> A -> Prop
Hh:forall P : A -> Prop, (forall x0 : A, P x0 <-> h P x0) /\ (forall Q : A -> Prop, (forall x0 : A, P x0 <-> Q x0) -> h P = h Q)
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class
x:A
Hx:forall x0 : A, class_of x x0 <-> h (class_of x) x0
Huniqx:forall Q : A -> Prop, (forall x0 : A, class_of x x0 <-> Q x0) -> h (class_of x) = h Q
y:A
Hxy:R x y

h (class_of x) = h (class_of y)
FunChoice:FunctionalChoice
PredExtRepr:ExtensionalPredicateRepresentative
PI:ProofIrrelevance
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
isclass:=fun P : A -> Prop => exists x0 : A, P x0:(A -> Prop) -> Prop
class:={P : A -> Prop | isclass P}:Type
contains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Prop
class_of:=fun a : A => R a:A -> A -> Prop
f:class -> A
Hf:forall x0 : class, contains x0 (f x0)
h:(A -> Prop) -> A -> Prop
Hh:forall P : A -> Prop, (forall x0 : A, P x0 <-> h P x0) /\ (forall Q : A -> Prop, (forall x0 : A, P x0 <-> Q x0) -> h P = h Q)
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class
x:A
Hx:forall x0 : A, class_of x x0 <-> h (class_of x) x0
Huniqx:forall Q : A -> Prop, (forall x0 : A, class_of x x0 <-> Q x0) -> h (class_of x) = h Q
y:A
Hxy:R x y

forall x0 : A, class_of x x0 <-> class_of y x0
FunChoice:FunctionalChoice
PredExtRepr:ExtensionalPredicateRepresentative
PI:ProofIrrelevance
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
isclass:=fun P : A -> Prop => exists x0 : A, P x0:(A -> Prop) -> Prop
class:={P : A -> Prop | isclass P}:Type
contains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Prop
class_of:=fun a : A => R a:A -> A -> Prop
f:class -> A
Hf:forall x0 : class, contains x0 (f x0)
h:(A -> Prop) -> A -> Prop
Hh:forall P : A -> Prop, (forall x0 : A, P x0 <-> h P x0) /\ (forall Q : A -> Prop, (forall x0 : A, P x0 <-> Q x0) -> h P = h Q)
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class
x:A
Hx:forall x0 : A, class_of x x0 <-> h (class_of x) x0
Huniqx:forall Q : A -> Prop, (forall x0 : A, class_of x x0 <-> Q x0) -> h (class_of x) = h Q
y:A
Hxy:R x y
z:A

class_of x z <-> class_of y z
FunChoice:FunctionalChoice
PredExtRepr:ExtensionalPredicateRepresentative
PI:ProofIrrelevance
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
isclass:=fun P : A -> Prop => exists x0 : A, P x0:(A -> Prop) -> Prop
class:={P : A -> Prop | isclass P}:Type
contains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Prop
class_of:=fun a : A => R a:A -> A -> Prop
f:class -> A
Hf:forall x0 : class, contains x0 (f x0)
h:(A -> Prop) -> A -> Prop
Hh:forall P : A -> Prop, (forall x0 : A, P x0 <-> h P x0) /\ (forall Q : A -> Prop, (forall x0 : A, P x0 <-> Q x0) -> h P = h Q)
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class
x:A
Hx:forall x0 : A, class_of x x0 <-> h (class_of x) x0
Huniqx:forall Q : A -> Prop, (forall x0 : A, class_of x x0 <-> Q x0) -> h (class_of x) = h Q
y:A
Hxy:R x y
z:A

R x z <-> R y z
firstorder.
FunChoice:FunctionalChoice
PredExtRepr:ExtensionalPredicateRepresentative
PI:ProofIrrelevance
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
isclass:=fun P : A -> Prop => exists x0 : A, P x0:(A -> Prop) -> Prop
class:={P : A -> Prop | isclass P}:Type
contains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Prop
class_of:=fun a : A => R a:A -> A -> Prop
f:class -> A
Hf:forall x0 : class, contains x0 (f x0)
h:(A -> Prop) -> A -> Prop
Hh:forall P : A -> Prop, (forall x0 : A, P x0 <-> h P x0) /\ (forall Q : A -> Prop, (forall x0 : A, P x0 <-> Q x0) -> h P = h Q)
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class
x:A
Hx:forall x0 : A, class_of x x0 <-> h (class_of x) x0
Huniqx:forall Q : A -> Prop, (forall x0 : A, class_of x x0 <-> Q x0) -> h (class_of x) = h Q
y:A
Hxy:R x y
Heq1:h (class_of x) = h (class_of y)

f' x = f' y
FunChoice:FunctionalChoice
PredExtRepr:ExtensionalPredicateRepresentative
PI:ProofIrrelevance
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
isclass:=fun P : A -> Prop => exists x0 : A, P x0:(A -> Prop) -> Prop
class:={P : A -> Prop | isclass P}:Type
contains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Prop
class_of:=fun a : A => R a:A -> A -> Prop
f:class -> A
Hf:forall x0 : class, contains x0 (f x0)
h:(A -> Prop) -> A -> Prop
Hh:forall P : A -> Prop, (forall x0 : A, P x0 <-> h P x0) /\ (forall Q : A -> Prop, (forall x0 : A, P x0 <-> Q x0) -> h P = h Q)
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class
x:A
Hx:forall x0 : A, class_of x x0 <-> h (class_of x) x0
Huniqx:forall Q : A -> Prop, (forall x0 : A, class_of x x0 <-> Q x0) -> h (class_of x) = h Q
y:A
Hxy:R x y
Heq1:h (class_of x) = h (class_of y)

rew [isclass] Heq1 in Hisclass x = Hisclass y
FunChoice:FunctionalChoice
PredExtRepr:ExtensionalPredicateRepresentative
PI:ProofIrrelevance
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
isclass:=fun P : A -> Prop => exists x0 : A, P x0:(A -> Prop) -> Prop
class:={P : A -> Prop | isclass P}:Type
contains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Prop
class_of:=fun a : A => R a:A -> A -> Prop
f:class -> A
Hf:forall x0 : class, contains x0 (f x0)
h:(A -> Prop) -> A -> Prop
Hh:forall P : A -> Prop, (forall x0 : A, P x0 <-> h P x0) /\ (forall Q : A -> Prop, (forall x0 : A, P x0 <-> Q x0) -> h P = h Q)
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class
x:A
Hx:forall x0 : A, class_of x x0 <-> h (class_of x) x0
Huniqx:forall Q : A -> Prop, (forall x0 : A, class_of x x0 <-> Q x0) -> h (class_of x) = h Q
y:A
Hxy:R x y
Heq1:h (class_of x) = h (class_of y)
Heq2:rew [isclass] Heq1 in Hisclass x = Hisclass y
f' x = f' y
FunChoice:FunctionalChoice
PredExtRepr:ExtensionalPredicateRepresentative
PI:ProofIrrelevance
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
isclass:=fun P : A -> Prop => exists x0 : A, P x0:(A -> Prop) -> Prop
class:={P : A -> Prop | isclass P}:Type
contains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Prop
class_of:=fun a : A => R a:A -> A -> Prop
f:class -> A
Hf:forall x0 : class, contains x0 (f x0)
h:(A -> Prop) -> A -> Prop
Hh:forall P : A -> Prop, (forall x0 : A, P x0 <-> h P x0) /\ (forall Q : A -> Prop, (forall x0 : A, P x0 <-> Q x0) -> h P = h Q)
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class
x:A
Hx:forall x0 : A, class_of x x0 <-> h (class_of x) x0
Huniqx:forall Q : A -> Prop, (forall x0 : A, class_of x x0 <-> Q x0) -> h (class_of x) = h Q
y:A
Hxy:R x y
Heq1:h (class_of x) = h (class_of y)

rew [isclass] Heq1 in Hisclass x = Hisclass y
apply PI.
FunChoice:FunctionalChoice
PredExtRepr:ExtensionalPredicateRepresentative
PI:ProofIrrelevance
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
isclass:=fun P : A -> Prop => exists x0 : A, P x0:(A -> Prop) -> Prop
class:={P : A -> Prop | isclass P}:Type
contains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Prop
class_of:=fun a : A => R a:A -> A -> Prop
f:class -> A
Hf:forall x0 : class, contains x0 (f x0)
h:(A -> Prop) -> A -> Prop
Hh:forall P : A -> Prop, (forall x0 : A, P x0 <-> h P x0) /\ (forall Q : A -> Prop, (forall x0 : A, P x0 <-> Q x0) -> h P = h Q)
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class
x:A
Hx:forall x0 : A, class_of x x0 <-> h (class_of x) x0
Huniqx:forall Q : A -> Prop, (forall x0 : A, class_of x x0 <-> Q x0) -> h (class_of x) = h Q
y:A
Hxy:R x y
Heq1:h (class_of x) = h (class_of y)
Heq2:rew [isclass] Heq1 in Hisclass x = Hisclass y

f' x = f' y
FunChoice:FunctionalChoice
PredExtRepr:ExtensionalPredicateRepresentative
PI:ProofIrrelevance
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
isclass:=fun P : A -> Prop => exists x0 : A, P x0:(A -> Prop) -> Prop
class:={P : A -> Prop | isclass P}:Type
contains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Prop
class_of:=fun a : A => R a:A -> A -> Prop
f:class -> A
Hf:forall x0 : class, contains x0 (f x0)
h:(A -> Prop) -> A -> Prop
Hh:forall P : A -> Prop, (forall x0 : A, P x0 <-> h P x0) /\ (forall Q : A -> Prop, (forall x0 : A, P x0 <-> Q x0) -> h P = h Q)
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class
x:A
Hx:forall x0 : A, class_of x x0 <-> h (class_of x) x0
Huniqx:forall Q : A -> Prop, (forall x0 : A, class_of x x0 <-> Q x0) -> h (class_of x) = h Q
y:A
Hxy:R x y
Heq1:h (class_of x) = h (class_of y)
Heq2:rew [isclass] Heq1 in Hisclass x = Hisclass y

exist isclass (h (class_of x)) (Hisclass x) = exist isclass (h (class_of y)) (Hisclass y)
FunChoice:FunctionalChoice
PredExtRepr:ExtensionalPredicateRepresentative
PI:ProofIrrelevance
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
isclass:=fun P : A -> Prop => exists x0 : A, P x0:(A -> Prop) -> Prop
class:={P : A -> Prop | isclass P}:Type
contains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Prop
class_of:=fun a : A => R a:A -> A -> Prop
f:class -> A
Hf:forall x0 : class, contains x0 (f x0)
h:(A -> Prop) -> A -> Prop
Hh:forall P : A -> Prop, (forall x0 : A, P x0 <-> h P x0) /\ (forall Q : A -> Prop, (forall x0 : A, P x0 <-> Q x0) -> h P = h Q)
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class
x:A
Hx:forall x0 : A, class_of x x0 <-> h (class_of x) x0
Huniqx:forall Q : A -> Prop, (forall x0 : A, class_of x x0 <-> Q x0) -> h (class_of x) = h Q
y:A
Hxy:R x y
Heq1:h (class_of x) = h (class_of y)
Heq2:rew [isclass] Heq1 in Hisclass x = Hisclass y

exist isclass (h (class_of x)) (Hisclass x) = exist isclass (h (class_of y)) (rew [isclass] Heq1 in Hisclass x)
FunChoice:FunctionalChoice
PredExtRepr:ExtensionalPredicateRepresentative
PI:ProofIrrelevance
A:Type
R:A -> A -> Prop
Hrefl:Reflexive R
Hsym:Symmetric R
Htrans:Transitive R
isclass:=fun P : A -> Prop => exists x0 : A, P x0:(A -> Prop) -> Prop
class:={P : A -> Prop | isclass P}:Type
contains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Prop
class_of:=fun a : A => R a:A -> A -> Prop
f:class -> A
Hf:forall x0 : class, contains x0 (f x0)
h:(A -> Prop) -> A -> Prop
Hh:forall P : A -> Prop, (forall x0 : A, P x0 <-> h P x0) /\ (forall Q : A -> Prop, (forall x0 : A, P x0 <-> Q x0) -> h P = h Q)
Hisclass:forall a : A, isclass (h (class_of a))
f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> class
x:A
Hx:forall x0 : A, class_of x x0 <-> h (class_of x) x0
Huniqx:forall Q : A -> Prop, (forall x0 : A, class_of x x0 <-> Q x0) -> h (class_of x) = h Q
y:A
Hxy:R x y
Heq1:h (class_of x) = h (class_of y)
Heq2:rew [isclass] Heq1 in Hisclass x = Hisclass y

exist isclass (h (class_of x)) (Hisclass x) = exist isclass (h (class_of x)) (rew [isclass] eq_refl in Hisclass x)
reflexivity. Qed.

FunctionalChoice /\ ExtensionalPredicateRepresentative /\ ProofIrrelevance <-> SetoidFunctionalChoice

FunctionalChoice /\ ExtensionalPredicateRepresentative /\ ProofIrrelevance <-> SetoidFunctionalChoice

FunctionalChoice /\ ExtensionalPredicateRepresentative /\ ProofIrrelevance -> SetoidFunctionalChoice

SetoidFunctionalChoice -> FunctionalChoice /\ ExtensionalPredicateRepresentative /\ ProofIrrelevance

FunctionalChoice /\ ExtensionalPredicateRepresentative /\ ProofIrrelevance -> SetoidFunctionalChoice
FunChoice:FunctionalChoice
ExtPredRepr:ExtensionalPredicateRepresentative
PI:ProofIrrelevance

SetoidFunctionalChoice
FunChoice:FunctionalChoice
ExtPredRepr:ExtensionalPredicateRepresentative
PI:ProofIrrelevance

FunctionalRelReification
FunChoice:FunctionalChoice
ExtPredRepr:ExtensionalPredicateRepresentative
PI:ProofIrrelevance
RepresentativeFunctionalChoice
FunChoice:FunctionalChoice
ExtPredRepr:ExtensionalPredicateRepresentative
PI:ProofIrrelevance

FunctionalRelReification
FunChoice:FunctionalChoice
ExtPredRepr:ExtensionalPredicateRepresentative
PI:ProofIrrelevance
A, B:Type

FunctionalRelReification_on A B
now apply fun_choice_imp_functional_rel_reification.
FunChoice:FunctionalChoice
ExtPredRepr:ExtensionalPredicateRepresentative
PI:ProofIrrelevance

RepresentativeFunctionalChoice
now apply fun_choice_and_ext_pred_ext_and_proof_irrel_imp_setoid_fun_choice.

SetoidFunctionalChoice -> FunctionalChoice /\ ExtensionalPredicateRepresentative /\ ProofIrrelevance
SetoidFunChoice:SetoidFunctionalChoice

FunctionalChoice /\ ExtensionalPredicateRepresentative /\ ProofIrrelevance
SetoidFunChoice:SetoidFunctionalChoice

FunctionalChoice
SetoidFunChoice:SetoidFunctionalChoice
ExtensionalPredicateRepresentative
SetoidFunChoice:SetoidFunctionalChoice
ProofIrrelevance
SetoidFunChoice:SetoidFunctionalChoice

FunctionalChoice
now intros A B; apply setoid_fun_choice_imp_fun_choice.
SetoidFunChoice:SetoidFunctionalChoice

ExtensionalPredicateRepresentative
SetoidFunChoice:SetoidFunctionalChoice

RepresentativeFunctionalChoice
now apply setoid_fun_choice_imp_repr_fun_choice.
SetoidFunChoice:SetoidFunctionalChoice

ProofIrrelevance
SetoidFunChoice:SetoidFunctionalChoice

forall (A : Prop) (a1 a2 : A), a1 = a2
SetoidFunChoice:SetoidFunctionalChoice

forall A : Prop, A \/ ~ A
SetoidFunChoice:SetoidFunctionalChoice

RepresentativeFunctionalChoice
now apply setoid_fun_choice_imp_repr_fun_choice. Qed. (**********************************************************************)

Compatibility notations

Notation description_rel_choice_imp_funct_choice :=
  functional_rel_reification_and_rel_choice_imp_fun_choice (only parsing).

Notation funct_choice_imp_rel_choice := fun_choice_imp_rel_choice (only parsing).

Notation FunChoice_Equiv_RelChoice_and_ParamDefinDescr :=
 fun_choice_iff_rel_choice_and_functional_rel_reification (only parsing).

Notation funct_choice_imp_description := fun_choice_imp_functional_rel_reification (only parsing).