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.
[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. (**********************************************************************)
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.
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 }.
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
- IPL_2 = 2nd-order impredicative minimal predicate logic (with ex. quant.)
- IPL^2 = 2nd-order functional minimal predicate logic (with ex. quant.)
- IPL_2^2 = 2nd-order impredicative, 2nd-order functional minimal pred. logic (with ex. quant.)
(**********************************************************************)
(* 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
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 Bforall A B : Type, FunctionalRelReification_on A B -> RelationalChoice_on A B -> FunctionalChoice_on A BA, B:TypeDescr:FunctionalRelReification_on A BRelCh:RelationalChoice_on A BR:A -> B -> PropH:forall x : A, exists y : B, R x yexists f : A -> B, forall x : A, R x (f x)A, B:TypeDescr:FunctionalRelReification_on A BRelCh:RelationalChoice_on A BR:A -> B -> PropH:forall x : A, exists y : B, R x yR':A -> B -> PropHR'R:subrelation R' RH0:forall x : A, exists ! y : B, R' x yexists f : A -> B, forall x : A, R x (f x)A, B:TypeDescr:FunctionalRelReification_on A BRelCh:RelationalChoice_on A BR:A -> B -> PropH:forall x : A, exists y : B, R x yR':A -> B -> PropHR'R:subrelation R' RH0:forall x : A, exists ! y : B, R' x yforall x : A, exists ! y : B, R' x yA, B:TypeDescr:FunctionalRelReification_on A BRelCh:RelationalChoice_on A BR:A -> B -> PropH:forall x : A, exists y : B, R x yR':A -> B -> PropHR'R:subrelation R' RH0:forall x : A, exists ! y : B, R' x yf:A -> BHf:forall x : A, R' x (f x)exists f0 : A -> B, forall x : A, R x (f0 x)A, B:TypeDescr:FunctionalRelReification_on A BRelCh:RelationalChoice_on A BR:A -> B -> PropH:forall x : A, exists y : B, R x yR':A -> B -> PropHR'R:subrelation R' RH0:forall x : A, exists ! y : B, R' x yf:A -> BHf:forall x : A, R' x (f x)exists f0 : A -> B, forall x : A, R x (f0 x)A, B:TypeDescr:FunctionalRelReification_on A BRelCh:RelationalChoice_on A BR:A -> B -> PropH:forall x0 : A, exists y : B, R x0 yR':A -> B -> PropHR'R:subrelation R' RH0:forall x0 : A, exists ! y : B, R' x0 yf:A -> BHf:forall x0 : A, R' x0 (f x0)x:AR x (f x)A, B:TypeDescr:FunctionalRelReification_on A BRelCh:RelationalChoice_on A BR:A -> B -> PropH:forall x0 : A, exists y0 : B, R x0 y0R':A -> B -> PropHR'R:subrelation R' RH0:forall x0 : A, exists ! y0 : B, R' x0 y0f:A -> BHf:forall x0 : A, R' x0 (f x0)x:Ay:BHR'xy:R' x yHuniq:forall x' : B, R' x x' -> y = x'R x (f x)apply HR'R; assumption. Qed.A, B:TypeDescr:FunctionalRelReification_on A BRelCh:RelationalChoice_on A BR:A -> B -> PropH:forall x0 : A, exists y0 : B, R x0 y0R':A -> B -> PropHR'R:subrelation R' RH0:forall x0 : A, exists ! y0 : B, R' x0 y0f:A -> BHf:forall x0 : A, R' x0 (f x0)x:Ay:BHR'xy:R' x yHuniq:forall x' : B, R' x x' -> y = x'R x yforall A B : Type, FunctionalChoice_on A B -> RelationalChoice_on A Bforall A B : Type, FunctionalChoice_on A B -> RelationalChoice_on A BA, B:TypeFunCh:FunctionalChoice_on A BR:A -> B -> PropH:forall x : A, exists y : B, R x yexists R' : A -> B -> Prop, subrelation R' R /\ (forall x : A, exists ! y : B, R' x y)A, B:TypeFunCh:FunctionalChoice_on A BR:A -> B -> PropH:forall x : A, exists y : B, R x yf:A -> BH0: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:TypeFunCh:FunctionalChoice_on A BR:A -> B -> PropH:forall x : A, exists y : B, R x yf:A -> BH0: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:TypeFunCh:FunctionalChoice_on A BR:A -> B -> PropH:forall x : A, exists y : B, R x yf:A -> BH0:forall x : A, R x (f x)subrelation (fun (x : A) (y : B) => f x = y) RA, B:TypeFunCh:FunctionalChoice_on A BR:A -> B -> PropH:forall x : A, exists y : B, R x yf:A -> BH0:forall x : A, R x (f x)forall x : A, exists ! y : B, f x = yA, B:TypeFunCh:FunctionalChoice_on A BR:A -> B -> PropH:forall x : A, exists y : B, R x yf:A -> BH0:forall x : A, R x (f x)forall x : A, exists ! y : B, f x = yA, B:TypeFunCh:FunctionalChoice_on A BR:A -> B -> PropH:forall x0 : A, exists y : B, R x0 yf:A -> BH0:forall x0 : A, R x0 (f x0)x:Af x = f xA, B:TypeFunCh:FunctionalChoice_on A BR:A -> B -> PropH:forall x0 : A, exists y : B, R x0 yf:A -> BH0:forall x0 : A, R x0 (f x0)x:Aforall x' : B, f x = x' -> f x = x'trivial. Qed.A, B:TypeFunCh:FunctionalChoice_on A BR:A -> B -> PropH:forall x0 : A, exists y : B, R x0 yf:A -> BH0:forall x0 : A, R x0 (f x0)x:Aforall x' : B, f x = x' -> f x = x'forall A B : Type, FunctionalChoice_on A B -> FunctionalRelReification_on A Bforall A B : Type, FunctionalChoice_on A B -> FunctionalRelReification_on A BA, B:TypeFunCh:FunctionalChoice_on A BR:A -> B -> PropH:forall x : A, exists ! y : B, R x yexists f : A -> B, forall x : A, R x (f x)(* 1 *)A, B:TypeFunCh:FunctionalChoice_on A BR:A -> B -> PropH:forall x : A, exists ! y : B, R x yforall x : A, exists y : B, R x yA, B:TypeFunCh:FunctionalChoice_on A BR:A -> B -> PropH:forall x : A, exists ! y : B, R x yf:A -> BH0:forall x : A, R x (f x)exists f0 : A -> B, forall x : A, R x (f0 x)A, B:TypeFunCh:FunctionalChoice_on A BR:A -> B -> PropH:forall x0 : A, exists ! y : B, R x0 yx:Aexists y : B, R x yA, B:TypeFunCh:FunctionalChoice_on A BR:A -> B -> PropH:forall x : A, exists ! y : B, R x yf:A -> BH0:forall x : A, R x (f x)exists f0 : A -> B, forall x : A, R x (f0 x)A, B:TypeFunCh:FunctionalChoice_on A BR:A -> B -> PropH:forall x0 : A, exists ! y0 : B, R x0 y0x:Ay:BHRxy:R x yexists y0 : B, R x y0A, B:TypeFunCh:FunctionalChoice_on A BR:A -> B -> PropH:forall x : A, exists ! y : B, R x yf:A -> BH0:forall x : A, R x (f x)exists f0 : A -> B, forall x : A, R x (f0 x)(* 2 *) exists f; exact H0. Qed.A, B:TypeFunCh:FunctionalChoice_on A BR:A -> B -> PropH:forall x : A, exists ! y : B, R x yf:A -> BH0:forall x : A, R x (f x)exists f0 : A -> B, forall x : A, R x (f0 x)forall A B : Type, FunctionalChoice_on A B <-> RelationalChoice_on A B /\ FunctionalRelReification_on A Bforall A B : Type, FunctionalChoice_on A B <-> RelationalChoice_on A B /\ FunctionalRelReification_on A BA, B:TypeFunctionalChoice_on A B <-> RelationalChoice_on A B /\ FunctionalRelReification_on A BA, B:TypeFunctionalChoice_on A B -> RelationalChoice_on A B /\ FunctionalRelReification_on A BA, B:TypeRelationalChoice_on A B /\ FunctionalRelReification_on A B -> FunctionalChoice_on A Bintros [H H0]; exact (functional_rel_reification_and_rel_choice_imp_fun_choice H0 H). Qed. (**********************************************************************)A, B:TypeRelationalChoice_on A B /\ FunctionalRelReification_on A B -> FunctionalChoice_on A B
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)
(**********************************************************************)
RelationalChoice -> ProofIrrelevance -> GuardedRelationalChoiceRelationalChoice -> ProofIrrelevance -> GuardedRelationalChoicerel_choice:RelationalChoiceproof_irrel:ProofIrrelevanceGuardedRelationalChoicerel_choice:RelationalChoiceproof_irrel:ProofIrrelevanceA, B:TypeP:A -> PropR:A -> B -> PropH:forall x : A, P x -> exists y : B, R x yexists R' : A -> B -> Prop, subrelation R' R /\ (forall x : A, P x -> exists ! y : B, R' x y)rel_choice:RelationalChoiceproof_irrel:ProofIrrelevanceA, B:TypeP:A -> PropR:A -> B -> PropH:forall x : A, P x -> exists y : B, R x yforall x : {x : A & P x}, exists y : B, R (projT1 x) yrel_choice:RelationalChoiceproof_irrel:ProofIrrelevanceA, B:TypeP:A -> PropR:A -> B -> PropH:forall x : A, P x -> exists y : B, R x yR':{x : A & P x} -> B -> PropHR'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 yexists R'0 : A -> B -> Prop, subrelation R'0 R /\ (forall x : A, P x -> exists ! y : B, R'0 x y)rel_choice:RelationalChoiceproof_irrel:ProofIrrelevanceA, B:TypeP:A -> PropR:A -> B -> PropH:forall x0 : A, P x0 -> exists y : B, R x0 yx:AHPx:P xexists y : B, R (projT1 (existT P x HPx)) yrel_choice:RelationalChoiceproof_irrel:ProofIrrelevanceA, B:TypeP:A -> PropR:A -> B -> PropH:forall x : A, P x -> exists y : B, R x yR':{x : A & P x} -> B -> PropHR'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 yexists R'0 : A -> B -> Prop, subrelation R'0 R /\ (forall x : A, P x -> exists ! y : B, R'0 x y)rel_choice:RelationalChoiceproof_irrel:ProofIrrelevanceA, B:TypeP:A -> PropR:A -> B -> PropH:forall x0 : A, P x0 -> exists y0 : B, R x0 y0x:AHPx:P xy:BHRxy:R x yexists y0 : B, R (projT1 (existT P x HPx)) y0rel_choice:RelationalChoiceproof_irrel:ProofIrrelevanceA, B:TypeP:A -> PropR:A -> B -> PropH:forall x : A, P x -> exists y : B, R x yR':{x : A & P x} -> B -> PropHR'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 yexists R'0 : A -> B -> Prop, subrelation R'0 R /\ (forall x : A, P x -> exists ! y : B, R'0 x y)rel_choice:RelationalChoiceproof_irrel:ProofIrrelevanceA, B:TypeP:A -> PropR:A -> B -> PropH:forall x : A, P x -> exists y : B, R x yR':{x : A & P x} -> B -> PropHR'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 yexists R'0 : A -> B -> Prop, subrelation R'0 R /\ (forall x : A, P x -> exists ! y : B, R'0 x y)rel_choice:RelationalChoiceproof_irrel:ProofIrrelevanceA, B:TypeP:A -> PropR:A -> B -> PropH:forall x : A, P x -> exists y : B, R x yR':{x : A & P x} -> B -> PropHR'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 yR'':=fun (x : A) (y : B) => exists H1 : P x, R' (existT P x H1) y:A -> B -> Propexists R'0 : A -> B -> Prop, subrelation R'0 R /\ (forall x : A, P x -> exists ! y : B, R'0 x y)rel_choice:RelationalChoiceproof_irrel:ProofIrrelevanceA, B:TypeP:A -> PropR:A -> B -> PropH:forall x : A, P x -> exists y : B, R x yR':{x : A & P x} -> B -> PropHR'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 yR'':=fun (x : A) (y : B) => exists H1 : P x, R' (existT P x H1) y:A -> B -> Propsubrelation R'' Rrel_choice:RelationalChoiceproof_irrel:ProofIrrelevanceA, B:TypeP:A -> PropR:A -> B -> PropH:forall x : A, P x -> exists y : B, R x yR':{x : A & P x} -> B -> PropHR'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 yR'':=fun (x : A) (y : B) => exists H1 : P x, R' (existT P x H1) y:A -> B -> Propforall x : A, P x -> exists ! y : B, R'' x yrel_choice:RelationalChoiceproof_irrel:ProofIrrelevanceA, B:TypeP:A -> PropR:A -> B -> PropH:forall x0 : A, P x0 -> exists y0 : B, R x0 y0R':{x : A & P x} -> B -> PropHR'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 y0R'':=fun (x0 : A) (y0 : B) => exists H1 : P x0, R' (existT P x0 H1) y0:A -> B -> Propx:Ay:BHPx:P xHR'xy:R' (existT P x HPx) yR x yrel_choice:RelationalChoiceproof_irrel:ProofIrrelevanceA, B:TypeP:A -> PropR:A -> B -> PropH:forall x : A, P x -> exists y : B, R x yR':{x : A & P x} -> B -> PropHR'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 yR'':=fun (x : A) (y : B) => exists H1 : P x, R' (existT P x H1) y:A -> B -> Propforall x : A, P x -> exists ! y : B, R'' x yrel_choice:RelationalChoiceproof_irrel:ProofIrrelevanceA, B:TypeP:A -> PropR:A -> B -> PropH:forall x : A, P x -> exists y : B, R x yR':{x : A & P x} -> B -> PropHR'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 yR'':=fun (x : A) (y : B) => exists H1 : P x, R' (existT P x H1) y:A -> B -> Propforall x : A, P x -> exists ! y : B, R'' x yrel_choice:RelationalChoiceproof_irrel:ProofIrrelevanceA, B:TypeP:A -> PropR:A -> B -> PropH:forall x0 : A, P x0 -> exists y : B, R x0 yR':{x : A & P x} -> B -> PropHR'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 yR'':=fun (x0 : A) (y : B) => exists H1 : P x0, R' (existT P x0 H1) y:A -> B -> Propx:AHPx:P xexists ! y : B, R'' x yrel_choice:RelationalChoiceproof_irrel:ProofIrrelevanceA, B:TypeP:A -> PropR:A -> B -> PropH:forall x0 : A, P x0 -> exists y0 : B, R x0 y0R':{x : A & P x} -> B -> PropHR'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 y0R'':=fun (x0 : A) (y0 : B) => exists H1 : P x0, R' (existT P x0 H1) y0:A -> B -> Propx:AHPx:P xy:BHR'xy:R' (existT P x HPx) yHuniq:forall x' : B, R' (existT P x HPx) x' -> y = x'exists ! y0 : B, R'' x y0rel_choice:RelationalChoiceproof_irrel:ProofIrrelevanceA, B:TypeP:A -> PropR:A -> B -> PropH:forall x0 : A, P x0 -> exists y0 : B, R x0 y0R':{x : A & P x} -> B -> PropHR'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 y0R'':=fun (x0 : A) (y0 : B) => exists H1 : P x0, R' (existT P x0 H1) y0:A -> B -> Propx:AHPx:P xy:BHR'xy:R' (existT P x HPx) yHuniq:forall x' : B, R' (existT P x HPx) x' -> y = x'R'' x yrel_choice:RelationalChoiceproof_irrel:ProofIrrelevanceA, B:TypeP:A -> PropR:A -> B -> PropH:forall x0 : A, P x0 -> exists y0 : B, R x0 y0R':{x : A & P x} -> B -> PropHR'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 y0R'':=fun (x0 : A) (y0 : B) => exists H1 : P x0, R' (existT P x0 H1) y0:A -> B -> Propx:AHPx:P xy:BHR'xy:R' (existT P x HPx) yHuniq:forall x' : B, R' (existT P x HPx) x' -> y = x'forall x' : B, R'' x x' -> y = x'rel_choice:RelationalChoiceproof_irrel:ProofIrrelevanceA, B:TypeP:A -> PropR:A -> B -> PropH:forall x0 : A, P x0 -> exists y0 : B, R x0 y0R':{x : A & P x} -> B -> PropHR'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 y0R'':=fun (x0 : A) (y0 : B) => exists H1 : P x0, R' (existT P x0 H1) y0:A -> B -> Propx:AHPx:P xy:BHR'xy:R' (existT P x HPx) yHuniq:forall x' : B, R' (existT P x HPx) x' -> y = x'forall x' : B, R'' x x' -> y = x'rel_choice:RelationalChoiceproof_irrel:ProofIrrelevanceA, B:TypeP:A -> PropR:A -> B -> PropH:forall x0 : A, P x0 -> exists y0 : B, R x0 y0R':{x : A & P x} -> B -> PropHR'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 y0R'':=fun (x0 : A) (y0 : B) => exists H1 : P x0, R' (existT P x0 H1) y0:A -> B -> Propx:AHPx:P xy:BHR'xy:R' (existT P x HPx) yHuniq:forall x' : B, R' (existT P x HPx) x' -> y = x'y':BH'Px:P xHR'xy':R' (existT P x H'Px) y'y = y'rewrite proof_irrel with (a1 := HPx) (a2 := H'Px); exact HR'xy'. Qed.rel_choice:RelationalChoiceproof_irrel:ProofIrrelevanceA, B:TypeP:A -> PropR:A -> B -> PropH:forall x0 : A, P x0 -> exists y0 : B, R x0 y0R':{x : A & P x} -> B -> PropHR'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 y0R'':=fun (x0 : A) (y0 : B) => exists H1 : P x0, R' (existT P x0 H1) y0:A -> B -> Propx:AHPx:P xy:BHR'xy:R' (existT P x HPx) yHuniq:forall x' : B, R' (existT P x HPx) x' -> y = x'y':BH'Px:P xHR'xy':R' (existT P x H'Px) y'R' (existT P x HPx) y'forall A B : Type, inhabited B -> RelationalChoice_on A B -> IndependenceOfGeneralPremises -> GuardedRelationalChoice_on A Bforall A B : Type, inhabited B -> RelationalChoice_on A B -> IndependenceOfGeneralPremises -> GuardedRelationalChoice_on A BA, B:TypeInh:inhabited BAC_rel:RelationalChoice_on A BIndPrem:IndependenceOfGeneralPremisesP:A -> PropR:A -> B -> PropH:forall x : A, P x -> exists y : B, R x yexists R' : A -> B -> Prop, subrelation R' R /\ (forall x : A, P x -> exists ! y : B, R' x y)A, B:TypeInh:inhabited BAC_rel:RelationalChoice_on A BIndPrem:IndependenceOfGeneralPremisesP:A -> PropR:A -> B -> PropH:forall x : A, P x -> exists y : B, R x yforall x : A, exists y : B, P x -> R x yA, B:TypeInh:inhabited BAC_rel:RelationalChoice_on A BIndPrem:IndependenceOfGeneralPremisesP:A -> PropR:A -> B -> PropH:forall x : A, P x -> exists y : B, R x yR':A -> B -> PropHR'R:subrelation R' (fun (x : A) (y : B) => P x -> R x y)H0:forall x : A, exists ! y : B, R' x yexists R'0 : A -> B -> Prop, subrelation R'0 R /\ (forall x : A, P x -> exists ! y : B, R'0 x y)A, B:TypeInh:inhabited BAC_rel:RelationalChoice_on A BIndPrem:IndependenceOfGeneralPremisesP:A -> PropR:A -> B -> PropH:forall x0 : A, P x0 -> exists y : B, R x0 yx:Aexists y : B, P x -> R x yA, B:TypeInh:inhabited BAC_rel:RelationalChoice_on A BIndPrem:IndependenceOfGeneralPremisesP:A -> PropR:A -> B -> PropH:forall x : A, P x -> exists y : B, R x yR':A -> B -> PropHR'R:subrelation R' (fun (x : A) (y : B) => P x -> R x y)H0:forall x : A, exists ! y : B, R' x yexists R'0 : A -> B -> Prop, subrelation R'0 R /\ (forall x : A, P x -> exists ! y : B, R'0 x y)A, B:TypeInh:inhabited BAC_rel:RelationalChoice_on A BIndPrem:IndependenceOfGeneralPremisesP:A -> PropR:A -> B -> PropH:forall x0 : A, P x0 -> exists y : B, R x0 yx:Ainhabited BA, B:TypeInh:inhabited BAC_rel:RelationalChoice_on A BIndPrem:IndependenceOfGeneralPremisesP:A -> PropR:A -> B -> PropH:forall x0 : A, P x0 -> exists y : B, R x0 yx:AP x -> exists x0 : B, R x x0A, B:TypeInh:inhabited BAC_rel:RelationalChoice_on A BIndPrem:IndependenceOfGeneralPremisesP:A -> PropR:A -> B -> PropH:forall x : A, P x -> exists y : B, R x yR':A -> B -> PropHR'R:subrelation R' (fun (x : A) (y : B) => P x -> R x y)H0:forall x : A, exists ! y : B, R' x yexists R'0 : A -> B -> Prop, subrelation R'0 R /\ (forall x : A, P x -> exists ! y : B, R'0 x y)A, B:TypeInh:inhabited BAC_rel:RelationalChoice_on A BIndPrem:IndependenceOfGeneralPremisesP:A -> PropR:A -> B -> PropH:forall x0 : A, P x0 -> exists y : B, R x0 yx:AP x -> exists x0 : B, R x x0A, B:TypeInh:inhabited BAC_rel:RelationalChoice_on A BIndPrem:IndependenceOfGeneralPremisesP:A -> PropR:A -> B -> PropH:forall x : A, P x -> exists y : B, R x yR':A -> B -> PropHR'R:subrelation R' (fun (x : A) (y : B) => P x -> R x y)H0:forall x : A, exists ! y : B, R' x yexists R'0 : A -> B -> Prop, subrelation R'0 R /\ (forall x : A, P x -> exists ! y : B, R'0 x y)A, B:TypeInh:inhabited BAC_rel:RelationalChoice_on A BIndPrem:IndependenceOfGeneralPremisesP:A -> PropR:A -> B -> PropH:forall x0 : A, P x0 -> exists y : B, R x0 yx:AHx:P xexists x0 : B, R x x0A, B:TypeInh:inhabited BAC_rel:RelationalChoice_on A BIndPrem:IndependenceOfGeneralPremisesP:A -> PropR:A -> B -> PropH:forall x : A, P x -> exists y : B, R x yR':A -> B -> PropHR'R:subrelation R' (fun (x : A) (y : B) => P x -> R x y)H0:forall x : A, exists ! y : B, R' x yexists R'0 : A -> B -> Prop, subrelation R'0 R /\ (forall x : A, P x -> exists ! y : B, R'0 x y)A, B:TypeInh:inhabited BAC_rel:RelationalChoice_on A BIndPrem:IndependenceOfGeneralPremisesP:A -> PropR:A -> B -> PropH:forall x : A, P x -> exists y : B, R x yR':A -> B -> PropHR'R:subrelation R' (fun (x : A) (y : B) => P x -> R x y)H0:forall x : A, exists ! y : B, R' x yexists R'0 : A -> B -> Prop, subrelation R'0 R /\ (forall x : A, P x -> exists ! y : B, R'0 x y)firstorder. Qed.A, B:TypeInh:inhabited BAC_rel:RelationalChoice_on A BIndPrem:IndependenceOfGeneralPremisesP:A -> PropR:A -> B -> PropH:forall x : A, P x -> exists y : B, R x yR':A -> B -> PropHR'R:subrelation R' (fun (x : A) (y : B) => P x -> R x y)H0:forall x : A, exists ! y : B, R' x ysubrelation (fun (x : A) (y : B) => P x /\ R' x y) R /\ (forall x : A, P x -> exists ! y : B, P x /\ R' x y)forall A B : Type, GuardedRelationalChoice_on A B -> RelationalChoice_on A Bforall A B : Type, GuardedRelationalChoice_on A B -> RelationalChoice_on A BA, B:TypeGAC_rel:GuardedRelationalChoice_on A BR:A -> B -> PropH:forall x : A, exists y : B, R x yexists R' : A -> B -> Prop, subrelation R' R /\ (forall x : A, exists ! y : B, R' x y)A, B:TypeGAC_rel:GuardedRelationalChoice_on A BR:A -> B -> PropH:forall x : A, exists y : B, R x yforall x : A, True -> exists y : B, R x yA, B:TypeGAC_rel:GuardedRelationalChoice_on A BR:A -> B -> PropH:forall x : A, exists y : B, R x yR':A -> B -> PropHR'R:subrelation R' RH0:forall x : A, True -> exists ! y : B, R' x yexists R'0 : A -> B -> Prop, subrelation R'0 R /\ (forall x : A, exists ! y : B, R'0 x y)exists R'; firstorder. Qed.A, B:TypeGAC_rel:GuardedRelationalChoice_on A BR:A -> B -> PropH:forall x : A, exists y : B, R x yR':A -> B -> PropHR'R:subrelation R' RH0:forall x : A, True -> exists ! y : B, R' x yexists R'0 : A -> B -> Prop, subrelation R'0 R /\ (forall x : A, exists ! y : B, R'0 x y)ProofIrrelevance -> GuardedRelationalChoice <-> RelationalChoiceintuition auto using guarded_rel_choice_imp_rel_choice, rel_choice_and_proof_irrel_imp_guarded_rel_choice. Qed.ProofIrrelevance -> GuardedRelationalChoice <-> RelationalChoice
OAC_rel = GAC_rel
GuardedRelationalChoice <-> OmniscientRelationalChoiceGuardedRelationalChoice <-> OmniscientRelationalChoiceGuardedRelationalChoice -> OmniscientRelationalChoiceOmniscientRelationalChoice -> GuardedRelationalChoiceGAC_rel:GuardedRelationalChoiceA, B:TypeR:A -> B -> Propexists R' : A -> B -> Prop, subrelation R' R /\ (forall x : A, (exists y : B, R x y) -> exists ! y : B, R' x y)OmniscientRelationalChoice -> GuardedRelationalChoiceOmniscientRelationalChoice -> GuardedRelationalChoicedestruct (OAC_rel A B R) as (f,Hf); exists f; firstorder. Qed. (**********************************************************************)OAC_rel:OmniscientRelationalChoiceA, B:TypeP:A -> PropR:A -> B -> PropH:forall x : A, P x -> exists y : B, R x yexists R' : A -> B -> Prop, subrelation R' R /\ (forall x : A, P x -> exists ! y : B, R' x y)
AC_fun + IGP = GAC_fun
GuardedFunctionalChoice -> IndependenceOfGeneralPremisesGuardedFunctionalChoice -> IndependenceOfGeneralPremisesGAC_fun:GuardedFunctionalChoiceA:TypeP:A -> PropQ:PropInh:inhabited AH:Q -> exists x : A, P xexists x : A, Q -> P xGAC_fun:GuardedFunctionalChoiceA:TypeP:A -> PropQ:PropInh:inhabited AH:Q -> exists x : A, P xunit -> Q -> exists y : A, P yGAC_fun:GuardedFunctionalChoiceA:TypeP:A -> PropQ:PropInh:inhabited AH:Q -> exists x : A, P xf:unit -> AHf:forall x : unit, Q -> P (f x)exists x : A, Q -> P xexists (f tt); auto. Qed.GAC_fun:GuardedFunctionalChoiceA:TypeP:A -> PropQ:PropInh:inhabited AH:Q -> exists x : A, P xf:unit -> AHf:forall x : unit, Q -> P (f x)exists x : A, Q -> P xGuardedFunctionalChoice -> FunctionalChoiceOnInhabitedSetGuardedFunctionalChoice -> FunctionalChoiceOnInhabitedSetGAC_fun:GuardedFunctionalChoiceA, B:TypeInh:inhabited BR:A -> B -> PropH:forall x : A, exists y : B, R x yexists f : A -> B, forall x : A, R x (f x)GAC_fun:GuardedFunctionalChoiceA, B:TypeInh:inhabited BR:A -> B -> PropH:forall x : A, exists y : B, R x yforall x : A, True -> exists y : B, R x yGAC_fun:GuardedFunctionalChoiceA, B:TypeInh:inhabited BR:A -> B -> PropH:forall x : A, exists y : B, R x yf:A -> BHf:forall x : A, True -> R x (f x)exists f0 : A -> B, forall x : A, R x (f0 x)exists f; auto. Qed.GAC_fun:GuardedFunctionalChoiceA, B:TypeInh:inhabited BR:A -> B -> PropH:forall x : A, exists y : B, R x yf:A -> BHf:forall x : A, True -> R x (f x)exists f0 : A -> B, forall x : A, R x (f0 x)FunctionalChoiceOnInhabitedSet -> IndependenceOfGeneralPremises -> GuardedFunctionalChoiceFunctionalChoiceOnInhabitedSet -> IndependenceOfGeneralPremises -> GuardedFunctionalChoiceAC_fun:FunctionalChoiceOnInhabitedSetIndPrem:IndependenceOfGeneralPremisesA, B:TypeP:A -> PropR:A -> B -> PropInh:inhabited BH:forall x : A, P x -> exists y : B, R x yexists f : A -> B, forall x : A, P x -> R x (f x)intro x; apply IndPrem; eauto. Qed.AC_fun:FunctionalChoiceOnInhabitedSetIndPrem:IndependenceOfGeneralPremisesA, B:TypeP:A -> PropR:A -> B -> PropInh:inhabited BH:forall x : A, P x -> exists y : B, R x yforall x : A, exists y : B, P x -> R x yFunctionalChoiceOnInhabitedSet /\ IndependenceOfGeneralPremises <-> GuardedFunctionalChoiceintuition 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.FunctionalChoiceOnInhabitedSet /\ IndependenceOfGeneralPremises <-> GuardedFunctionalChoice
AC_fun + Drinker = OAC_fun
This was already observed by Bell [Bell]
OmniscientFunctionalChoice -> SmallDrinker'sParadoxOmniscientFunctionalChoice -> SmallDrinker'sParadoxOAC_fun:OmniscientFunctionalChoiceA:TypeP:A -> PropInh:inhabited Aexists x : A, (exists x0 : A, P x0) -> P xOAC_fun:OmniscientFunctionalChoiceA:TypeP:A -> PropInh:inhabited Ainhabited AOAC_fun:OmniscientFunctionalChoiceA:TypeP:A -> PropInh:inhabited Af:unit -> AHf:forall x : unit, (exists y : A, P y) -> P (f x)exists x : A, (exists x0 : A, P x0) -> P xexists (f tt); firstorder. Qed.OAC_fun:OmniscientFunctionalChoiceA:TypeP:A -> PropInh:inhabited Af:unit -> AHf:forall x : unit, (exists y : A, P y) -> P (f x)exists x : A, (exists x0 : A, P x0) -> P xOmniscientFunctionalChoice -> FunctionalChoiceOnInhabitedSetOmniscientFunctionalChoice -> FunctionalChoiceOnInhabitedSetOAC_fun:OmniscientFunctionalChoiceA, B:TypeInh:inhabited BR:A -> B -> PropH:forall x : A, exists y : B, R x yexists f : A -> B, forall x : A, R x (f x)exists f; firstorder. Qed.OAC_fun:OmniscientFunctionalChoiceA, B:TypeInh:inhabited BR:A -> B -> PropH:forall x : A, exists y : B, R x yf:A -> BHf:forall x : A, (exists y : B, R x y) -> R x (f x)exists f0 : A -> B, forall x : A, R x (f0 x)FunctionalChoiceOnInhabitedSet -> SmallDrinker'sParadox -> OmniscientFunctionalChoiceFunctionalChoiceOnInhabitedSet -> SmallDrinker'sParadox -> OmniscientFunctionalChoiceAC_fun:FunctionalChoiceOnInhabitedSetDrinker:SmallDrinker'sParadoxA, B:TypeR:A -> B -> PropInh:inhabited Bexists f : A -> B, forall x : A, (exists y : B, R x y) -> R x (f x)AC_fun:FunctionalChoiceOnInhabitedSetDrinker:SmallDrinker'sParadoxA, B:TypeR:A -> B -> PropInh:inhabited Bforall x : A, exists y : B, (exists y0 : B, R x y0) -> R x yAC_fun:FunctionalChoiceOnInhabitedSetDrinker:SmallDrinker'sParadoxA, B:TypeR:A -> B -> PropInh:inhabited Bf:A -> BHf: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.AC_fun:FunctionalChoiceOnInhabitedSetDrinker:SmallDrinker'sParadoxA, B:TypeR:A -> B -> PropInh:inhabited Bf:A -> BHf: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)FunctionalChoiceOnInhabitedSet /\ SmallDrinker'sParadox <-> OmniscientFunctionalChoiceintuition auto using omniscient_fun_choice_imp_small_drinker, omniscient_fun_choice_imp_fun_choice, fun_choice_and_small_drinker_imp_omniscient_fun_choice. Qed.FunctionalChoiceOnInhabitedSet /\ SmallDrinker'sParadox <-> OmniscientFunctionalChoice
OAC_fun = GAC_fun
This is derivable from the intuitionistic equivalence between IGP and Drinker
but we give a direct proof
GuardedFunctionalChoice <-> OmniscientFunctionalChoiceGuardedFunctionalChoice <-> OmniscientFunctionalChoiceGuardedFunctionalChoice -> OmniscientFunctionalChoiceOmniscientFunctionalChoice -> GuardedFunctionalChoiceGAC_fun:GuardedFunctionalChoiceA, B:TypeR:A -> B -> PropInh:inhabited Bexists f : A -> B, forall x : A, (exists y : B, R x y) -> R x (f x)OmniscientFunctionalChoice -> GuardedFunctionalChoiceOmniscientFunctionalChoice -> GuardedFunctionalChoiceOAC_fun:OmniscientFunctionalChoiceA, B:TypeP:A -> PropR:A -> B -> PropInh:inhabited BH:forall x : A, P x -> exists y : B, R x yexists f : A -> B, forall x : A, P x -> R x (f x)exists f; firstorder. Qed. (**********************************************************************)OAC_fun:OmniscientFunctionalChoiceA, B:TypeP:A -> PropR:A -> B -> PropInh:inhabited BH:forall x : A, P x -> exists y : B, R x yf:A -> BHf: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)
D_iota -> ID_iota
IotaStatement -> ConstructiveDefiniteDescriptionIotaStatement -> ConstructiveDefiniteDescriptionD_iota:IotaStatementA:TypeP:A -> PropH:exists ! x : A, P x{x : A | P x}D_iota:IotaStatementA:TypeP:A -> PropH:exists ! x : A, P xinhabited AD_iota:IotaStatementA:TypeP:A -> PropH:exists ! x0 : A, P x0x:AH1:(exists ! x0 : A, P x0) -> P x{x0 : A | P x0}exists x; apply H1; assumption. Qed.D_iota:IotaStatementA:TypeP:A -> PropH:exists ! x0 : A, P x0x:AH1:(exists ! x0 : A, P x0) -> P x{x0 : A | P x0}
ID_epsilon + Drinker <-> D_epsilon
EpsilonStatement -> ConstructiveIndefiniteDescriptionEpsilonStatement -> ConstructiveIndefiniteDescriptionD_epsilon:EpsilonStatementA:TypeP:A -> PropH:exists x : A, P x{x : A | P x}D_epsilon:EpsilonStatementA:TypeP:A -> PropH:exists x : A, P xinhabited AD_epsilon:EpsilonStatementA:TypeP:A -> PropH:exists x0 : A, P x0x:AH1:(exists x0 : A, P x0) -> P x{x0 : A | P x0}exists x; apply H1; assumption. Qed.D_epsilon:EpsilonStatementA:TypeP:A -> PropH:exists x0 : A, P x0x:AH1:(exists x0 : A, P x0) -> P x{x0 : A | P x0}SmallDrinker'sParadox -> ConstructiveIndefiniteDescription -> EpsilonStatementintros Drinkers D_epsilon A P Inh; apply D_epsilon; apply Drinkers; assumption. Qed.SmallDrinker'sParadox -> ConstructiveIndefiniteDescription -> EpsilonStatementEpsilonStatement -> SmallDrinker'sParadoxintros D_epsilon A P Inh; edestruct D_epsilon; eauto. Qed.EpsilonStatement -> SmallDrinker'sParadox((SmallDrinker'sParadox * ConstructiveIndefiniteDescription -> EpsilonStatement) * (EpsilonStatement -> SmallDrinker'sParadox * ConstructiveIndefiniteDescription))%typeintuition auto using epsilon_imp_constructive_indefinite_description, constructive_indefinite_description_and_small_drinker_imp_epsilon, epsilon_imp_small_drinker. Qed. (**********************************************************************)((SmallDrinker'sParadox * ConstructiveIndefiniteDescription -> EpsilonStatement) * (EpsilonStatement -> SmallDrinker'sParadox * ConstructiveIndefiniteDescription))%type
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 Rforall A : Type, FunctionalRelReification_on A nat -> forall R : A -> nat -> Prop, (forall (x : A) (y : nat), decidable (R x y)) -> FunctionalChoice_on_rel RA:TypeDescr:FunctionalRelReification_on A natforall R : A -> nat -> Prop, (forall (x : A) (y : nat), decidable (R x y)) -> FunctionalChoice_on_rel RA:TypeDescr:FunctionalRelReification_on A natR:A -> nat -> PropRdec:forall (x : A) (y : nat), decidable (R x y)H:forall x : A, exists y : nat, R x yexists f : A -> nat, forall x : A, R x (f x)A:TypeDescr:FunctionalRelReification_on A natR:A -> nat -> PropRdec:forall (x : A) (y : nat), decidable (R x y)H:forall x : A, exists y : nat, R x yR':=fun (x : A) (y : nat) => R x y /\ (forall y' : nat, R x y' -> y <= y'):A -> nat -> Propexists f : A -> nat, forall x : A, R x (f x)A:TypeDescr:FunctionalRelReification_on A natR:A -> nat -> PropRdec:forall (x : A) (y : nat), decidable (R x y)H:forall x : A, exists y : nat, R x yR':=fun (x : A) (y : nat) => R x y /\ (forall y' : nat, R x y' -> y <= y'):A -> nat -> Propforall x : A, exists ! y : nat, R' x yA:TypeDescr:FunctionalRelReification_on A natR:A -> nat -> PropRdec:forall (x : A) (y : nat), decidable (R x y)H:forall x : A, exists y : nat, R x yR':=fun (x : A) (y : nat) => R x y /\ (forall y' : nat, R x y' -> y <= y'):A -> nat -> Propf:A -> natHf:forall x : A, R' x (f x)exists f0 : A -> nat, forall x : A, R x (f0 x)A:TypeDescr:FunctionalRelReification_on A natR:A -> nat -> PropRdec:forall (x0 : A) (y : nat), decidable (R x0 y)H:forall x0 : A, exists y : nat, R x0 yR':=fun (x0 : A) (y : nat) => R x0 y /\ (forall y' : nat, R x0 y' -> y <= y'):A -> nat -> Propx:Aexists ! y : nat, R' x yA:TypeDescr:FunctionalRelReification_on A natR:A -> nat -> PropRdec:forall (x : A) (y : nat), decidable (R x y)H:forall x : A, exists y : nat, R x yR':=fun (x : A) (y : nat) => R x y /\ (forall y' : nat, R x y' -> y <= y'):A -> nat -> Propf:A -> natHf:forall x : A, R' x (f x)exists f0 : A -> nat, forall x : A, R x (f0 x)A:TypeDescr:FunctionalRelReification_on A natR:A -> nat -> PropRdec:forall (x0 : A) (y : nat), decidable (R x0 y)H:forall x0 : A, exists y : nat, R x0 yR':=fun (x0 : A) (y : nat) => R x0 y /\ (forall y' : nat, R x0 y' -> y <= y'):A -> nat -> Propx:Aforall n : nat, R x n \/ ~ R x nA:TypeDescr:FunctionalRelReification_on A natR:A -> nat -> PropRdec:forall (x0 : A) (y : nat), decidable (R x0 y)H:forall x0 : A, exists y : nat, R x0 yR':=fun (x0 : A) (y : nat) => R x0 y /\ (forall y' : nat, R x0 y' -> y <= y'):A -> nat -> Propx:Aexists n : nat, R x nA:TypeDescr:FunctionalRelReification_on A natR:A -> nat -> PropRdec:forall (x : A) (y : nat), decidable (R x y)H:forall x : A, exists y : nat, R x yR':=fun (x : A) (y : nat) => R x y /\ (forall y' : nat, R x y' -> y <= y'):A -> nat -> Propf:A -> natHf:forall x : A, R' x (f x)exists f0 : A -> nat, forall x : A, R x (f0 x)A:TypeDescr:FunctionalRelReification_on A natR:A -> nat -> PropRdec:forall (x0 : A) (y : nat), decidable (R x0 y)H:forall x0 : A, exists y : nat, R x0 yR':=fun (x0 : A) (y : nat) => R x0 y /\ (forall y' : nat, R x0 y' -> y <= y'):A -> nat -> Propx:Aexists n : nat, R x nA:TypeDescr:FunctionalRelReification_on A natR:A -> nat -> PropRdec:forall (x : A) (y : nat), decidable (R x y)H:forall x : A, exists y : nat, R x yR':=fun (x : A) (y : nat) => R x y /\ (forall y' : nat, R x y' -> y <= y'):A -> nat -> Propf:A -> natHf:forall x : A, R' x (f x)exists f0 : A -> nat, forall x : A, R x (f0 x)A:TypeDescr:FunctionalRelReification_on A natR:A -> nat -> PropRdec:forall (x : A) (y : nat), decidable (R x y)H:forall x : A, exists y : nat, R x yR':=fun (x : A) (y : nat) => R x y /\ (forall y' : nat, R x y' -> y <= y'):A -> nat -> Propf:A -> natHf:forall x : A, R' x (f x)exists f0 : A -> nat, forall x : A, R x (f0 x)A:TypeDescr:FunctionalRelReification_on A natR:A -> nat -> PropRdec:forall (x : A) (y : nat), decidable (R x y)H:forall x : A, exists y : nat, R x yR':=fun (x : A) (y : nat) => R x y /\ (forall y' : nat, R x y' -> y <= y'):A -> nat -> Propf:A -> natHf:forall x : A, R' x (f x)forall x : A, R x (f x)A:TypeDescr:FunctionalRelReification_on A natR:A -> nat -> PropRdec:forall (x0 : A) (y : nat), decidable (R x0 y)H:forall x0 : A, exists y : nat, R x0 yR':=fun (x0 : A) (y : nat) => R x0 y /\ (forall y' : nat, R x0 y' -> y <= y'):A -> nat -> Propf:A -> natHf:forall x0 : A, R' x0 (f x0)x:AR x (f x)assumption. Qed. (**********************************************************************)A:TypeDescr:FunctionalRelReification_on A natR:A -> nat -> PropRdec:forall (x0 : A) (y : nat), decidable (R x0 y)H:forall x0 : A, exists y : nat, R x0 yR':=fun (x0 : A) (y : nat) => R x0 y /\ (forall y' : nat, R x0 y' -> y <= y'):A -> nat -> Propf:A -> natHf:forall x0 : A, R' x0 (f x0)x:AHfx:R x (f x)R x (f x)
The easy part
DependentFunctionalChoice -> FunctionalChoiceDependentFunctionalChoice -> FunctionalChoiceAC_depfun:DependentFunctionalChoiceA, B:TypeR:A -> B -> PropH:forall x : A, exists y : B, R x yexists f : A -> B, forall x : A, R x (f x)exists f; trivial. Qed.AC_depfun:DependentFunctionalChoiceA, B:TypeR:A -> B -> PropH:forall x : A, exists y : B, R x yf:A -> BHf:forall x : A, R x (f x)exists f0 : A -> B, forall x : A, R x (f0 x)
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 -> DependentFunctionalChoiceFunctionalChoice -> DependentFunctionalChoiceAC_fun:FunctionalChoiceA:TypeB:A -> TypeR:forall x : A, B x -> PropH:forall x : A, exists y : B x, R x yexists f : forall x : A, B x, forall x : A, R x (f x)AC_fun:FunctionalChoiceA:TypeB:A -> TypeR:forall x : A, B x -> PropH:forall x : A, exists y : B x, R x yB':={x : A & B x}:Typeexists f : forall x : A, B x, forall x : A, R x (f x)AC_fun:FunctionalChoiceA:TypeB:A -> TypeR:forall x : A, B x -> PropH:forall x : A, exists y : B x, R x yB':={x : A & B x}:TypeR':=fun (x : A) (y : B') => projT1 y = x /\ R (projT1 y) (projT2 y):A -> B' -> Propexists f : forall x : A, B x, forall x : A, R x (f x)AC_fun:FunctionalChoiceA:TypeB:A -> TypeR:forall x : A, B x -> PropH:forall x : A, exists y : B x, R x yB':={x : A & B x}:TypeR':=fun (x : A) (y : B') => projT1 y = x /\ R (projT1 y) (projT2 y):A -> B' -> Propforall x : A, exists y : B', R' x yAC_fun:FunctionalChoiceA:TypeB:A -> TypeR:forall x : A, B x -> PropH:forall x : A, exists y : B x, R x yB':={x : A & B x}:TypeR':=fun (x : A) (y : B') => projT1 y = x /\ R (projT1 y) (projT2 y):A -> B' -> Propf: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:FunctionalChoiceA:TypeB:A -> TypeR:forall x0 : A, B x0 -> PropH:forall x0 : A, exists y : B x0, R x0 yB':={x0 : A & B x0}:TypeR':=fun (x0 : A) (y : B') => projT1 y = x0 /\ R (projT1 y) (projT2 y):A -> B' -> Propx:Aexists y : B', R' x yAC_fun:FunctionalChoiceA:TypeB:A -> TypeR:forall x : A, B x -> PropH:forall x : A, exists y : B x, R x yB':={x : A & B x}:TypeR':=fun (x : A) (y : B') => projT1 y = x /\ R (projT1 y) (projT2 y):A -> B' -> Propf: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:FunctionalChoiceA:TypeB:A -> TypeR:forall x0 : A, B x0 -> PropH:forall x0 : A, exists y0 : B x0, R x0 y0B':={x0 : A & B x0}:TypeR':=fun (x0 : A) (y0 : B') => projT1 y0 = x0 /\ R (projT1 y0) (projT2 y0):A -> B' -> Propx:Ay:B xHy:R x yexists y0 : B', R' x y0AC_fun:FunctionalChoiceA:TypeB:A -> TypeR:forall x : A, B x -> PropH:forall x : A, exists y : B x, R x yB':={x : A & B x}:TypeR':=fun (x : A) (y : B') => projT1 y = x /\ R (projT1 y) (projT2 y):A -> B' -> Propf: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:FunctionalChoiceA:TypeB:A -> TypeR:forall x0 : A, B x0 -> PropH:forall x0 : A, exists y0 : B x0, R x0 y0B':={x0 : A & B x0}:TypeR':=fun (x0 : A) (y0 : B') => projT1 y0 = x0 /\ R (projT1 y0) (projT2 y0):A -> B' -> Propx:Ay:B xHy:R x yR' x (existT (fun x0 : A => B x0) x y)AC_fun:FunctionalChoiceA:TypeB:A -> TypeR:forall x : A, B x -> PropH:forall x : A, exists y : B x, R x yB':={x : A & B x}:TypeR':=fun (x : A) (y : B') => projT1 y = x /\ R (projT1 y) (projT2 y):A -> B' -> Propf: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:FunctionalChoiceA:TypeB:A -> TypeR:forall x : A, B x -> PropH:forall x : A, exists y : B x, R x yB':={x : A & B x}:TypeR':=fun (x : A) (y : B') => projT1 y = x /\ R (projT1 y) (projT2 y):A -> B' -> Propf: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:FunctionalChoiceA:TypeB:A -> TypeR:forall x : A, B x -> PropH:forall x : A, exists y : B x, R x yB':={x : A & B x}:TypeR':=fun (x : A) (y : B') => projT1 y = x /\ R (projT1 y) (projT2 y):A -> B' -> Propf: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:FunctionalChoiceA:TypeB:A -> TypeR:forall x0 : A, B x0 -> PropH:forall x0 : A, exists y : B x0, R x0 yB':={x0 : A & B x0}:TypeR':=fun (x0 : A) (y : B') => projT1 y = x0 /\ R (projT1 y) (projT2 y):A -> B' -> Propf:A -> B'Hf:forall x0 : A, R' x0 (f x0)x:AHeq:projT1 (f x) = xHR: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)))destruct Heq using eq_indd; trivial. Qed.AC_fun:FunctionalChoiceA:TypeB:A -> TypeR:forall x1 : A, B x1 -> PropH:forall x1 : A, exists y : B x1, R x1 yB':={x1 : A & B x1}:TypeR':=fun (x1 : A) (y : B') => projT1 y = x1 /\ R (projT1 y) (projT2 y):A -> B' -> Propf:A -> B'Hf:forall x1 : A, R' x1 (f x1)x, x0:Ab:B x0Heq:x0 = xHR:R x0 bR x (eq_rect x0 (fun x1 : A => B x1) b x Heq)
FunctionalChoice -> InhabitedForallCommuteFunctionalChoice -> InhabitedForallCommutechoose0:FunctionalChoiceA:TypeB:A -> TypeHinhab:forall x : A, inhabited (B x)inhabited (forall x : A, B x)A:TypeB:A -> TypeHinhab:forall x : A, inhabited (B x)choose:DependentFunctionalChoiceinhabited (forall x : A, B x)A:TypeB:A -> TypeHinhab:forall x : A, inhabited (B x)choose:DependentFunctionalChoiceforall x : A, exists _ : B x, TrueA:TypeB:A -> TypeHinhab:forall x : A, inhabited (B x)choose:DependentFunctionalChoiceHexists:forall x : A, exists _ : B x, Trueinhabited (forall x : A, B x)A:TypeB:A -> TypeHinhab:forall x : A, inhabited (B x)choose:DependentFunctionalChoiceforall x : A, exists _ : B x, TrueA:TypeB:A -> TypeHinhab:forall x0 : A, inhabited (B x0)choose:DependentFunctionalChoicex:Ainhabited {_ : B x | True}intros y;exists y;exact I.A:TypeB:A -> TypeHinhab:forall x0 : A, inhabited (B x0)choose:DependentFunctionalChoicex:AB x -> {_ : B x | True}A:TypeB:A -> TypeHinhab:forall x : A, inhabited (B x)choose:DependentFunctionalChoiceHexists:forall x : A, exists _ : B x, Trueinhabited (forall x : A, B x)A:TypeB:A -> TypeHinhab:forall x : A, inhabited (B x)choose:DependentFunctionalChoiceHexists:exists _ : forall x : A, B x, A -> Trueinhabited (forall x : A, B x)exact (inhabits f). Qed.A:TypeB:A -> TypeHinhab:forall x : A, inhabited (B x)choose:DependentFunctionalChoicef:forall x : A, B xinhabited (forall x : A, B x)InhabitedForallCommute -> FunctionalChoiceInhabitedForallCommute -> FunctionalChoicechoose:InhabitedForallCommuteA, B:TypeR:A -> B -> PropHexists:forall x : A, exists y : B, R x yexists f : A -> B, forall x : A, R x (f x)choose:InhabitedForallCommuteA, B:TypeR:A -> B -> PropHexists:forall x : A, exists y : B, R x yforall x : A, inhabited {y : B | R x y}choose:InhabitedForallCommuteA, B:TypeR:A -> B -> PropHexists:forall x : A, exists y : B, R x yHinhab:forall x : A, inhabited {y : B | R x y}exists f : A -> B, forall x : A, R x (f x)intros x;apply exists_to_inhabited_sig;trivial.choose:InhabitedForallCommuteA, B:TypeR:A -> B -> PropHexists:forall x : A, exists y : B, R x yforall x : A, inhabited {y : B | R x y}choose:InhabitedForallCommuteA, B:TypeR:A -> B -> PropHexists:forall x : A, exists y : B, R x yHinhab:forall x : A, inhabited {y : B | R x y}exists f : A -> B, forall x : A, R x (f x)choose:InhabitedForallCommuteA, B:TypeR:A -> B -> PropHexists:forall x : A, exists y : B, R x yHinhab:inhabited (forall x : A, {y : B | R x y})exists f : A -> B, forall x : A, R x (f x)choose:InhabitedForallCommuteA, B:TypeR:A -> B -> PropHexists:forall x : A, exists y : B, R x yf:forall x : A, {y : B | R x y}exists f0 : A -> B, forall x : A, R x (f0 x)exact (fun x => proj2_sig (f x)). Qed.choose:InhabitedForallCommuteA, B:TypeR:A -> B -> PropHexists:forall x : A, exists y : B, R x yf:forall x : A, {y : B | R x y}forall x : A, R x (proj1_sig (f x))
The easy part
DependentFunctionalRelReification -> FunctionalRelReificationDependentFunctionalRelReification -> FunctionalRelReificationDepFunReify:DependentFunctionalRelReificationA, B:TypeR:A -> B -> PropH:forall x : A, exists ! y : B, R x yexists f : A -> B, forall x : A, R x (f x)exists f; trivial. Qed.DepFunReify:DependentFunctionalRelReificationA, B:TypeR:A -> B -> PropH:forall x : A, exists ! y : B, R x yf:A -> BHf:forall x : A, R x (f x)exists f0 : A -> B, forall x : A, R x (f0 x)
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 -> DependentFunctionalRelReificationFunctionalRelReification -> DependentFunctionalRelReificationAC_fun:FunctionalRelReificationA:TypeB:A -> TypeR:forall x : A, B x -> PropH:forall x : A, exists ! y : B x, R x yexists f : forall x : A, B x, forall x : A, R x (f x)AC_fun:FunctionalRelReificationA:TypeB:A -> TypeR:forall x : A, B x -> PropH:forall x : A, exists ! y : B x, R x yB':={x : A & B x}:Typeexists f : forall x : A, B x, forall x : A, R x (f x)AC_fun:FunctionalRelReificationA:TypeB:A -> TypeR:forall x : A, B x -> PropH:forall x : A, exists ! y : B x, R x yB':={x : A & B x}:TypeR':=fun (x : A) (y : B') => projT1 y = x /\ R (projT1 y) (projT2 y):A -> B' -> Propexists f : forall x : A, B x, forall x : A, R x (f x)AC_fun:FunctionalRelReificationA:TypeB:A -> TypeR:forall x : A, B x -> PropH:forall x : A, exists ! y : B x, R x yB':={x : A & B x}:TypeR':=fun (x : A) (y : B') => projT1 y = x /\ R (projT1 y) (projT2 y):A -> B' -> Propforall x : A, exists ! y : B', R' x yAC_fun:FunctionalRelReificationA:TypeB:A -> TypeR:forall x : A, B x -> PropH:forall x : A, exists ! y : B x, R x yB':={x : A & B x}:TypeR':=fun (x : A) (y : B') => projT1 y = x /\ R (projT1 y) (projT2 y):A -> B' -> Propf: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:FunctionalRelReificationA:TypeB:A -> TypeR:forall x0 : A, B x0 -> PropH:forall x0 : A, exists ! y : B x0, R x0 yB':={x0 : A & B x0}:TypeR':=fun (x0 : A) (y : B') => projT1 y = x0 /\ R (projT1 y) (projT2 y):A -> B' -> Propx:Aexists ! y : B', R' x yAC_fun:FunctionalRelReificationA:TypeB:A -> TypeR:forall x : A, B x -> PropH:forall x : A, exists ! y : B x, R x yB':={x : A & B x}:TypeR':=fun (x : A) (y : B') => projT1 y = x /\ R (projT1 y) (projT2 y):A -> B' -> Propf: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:FunctionalRelReificationA:TypeB:A -> TypeR:forall x0 : A, B x0 -> PropH:forall x0 : A, exists ! y0 : B x0, R x0 y0B':={x0 : A & B x0}:TypeR':=fun (x0 : A) (y0 : B') => projT1 y0 = x0 /\ R (projT1 y0) (projT2 y0):A -> B' -> Propx:Ay:B xHy:R x yHuni:forall x' : B x, R x x' -> y = x'exists ! y0 : B', R' x y0AC_fun:FunctionalRelReificationA:TypeB:A -> TypeR:forall x : A, B x -> PropH:forall x : A, exists ! y : B x, R x yB':={x : A & B x}:TypeR':=fun (x : A) (y : B') => projT1 y = x /\ R (projT1 y) (projT2 y):A -> B' -> Propf: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:FunctionalRelReificationA:TypeB:A -> TypeR:forall x0 : A, B x0 -> PropH:forall x0 : A, exists ! y0 : B x0, R x0 y0B':={x0 : A & B x0}:TypeR':=fun (x0 : A) (y0 : B') => projT1 y0 = x0 /\ R (projT1 y0) (projT2 y0):A -> B' -> Propx:Ay:B xHy:R x yHuni: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:FunctionalRelReificationA:TypeB:A -> TypeR:forall x : A, B x -> PropH:forall x : A, exists ! y : B x, R x yB':={x : A & B x}:TypeR':=fun (x : A) (y : B') => projT1 y = x /\ R (projT1 y) (projT2 y):A -> B' -> Propf: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:FunctionalRelReificationA:TypeB:A -> TypeR:forall x0 : A, B x0 -> PropH:forall x0 : A, exists ! y0 : B x0, R x0 y0B':={x0 : A & B x0}:TypeR':=fun (x0 : A) (y0 : B') => projT1 y0 = x0 /\ R (projT1 y0) (projT2 y0):A -> B' -> Propx:Ay:B xHy:R x yHuni: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:FunctionalRelReificationA:TypeB:A -> TypeR:forall x : A, B x -> PropH:forall x : A, exists ! y : B x, R x yB':={x : A & B x}:TypeR':=fun (x : A) (y : B') => projT1 y = x /\ R (projT1 y) (projT2 y):A -> B' -> Propf: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:FunctionalRelReificationA:TypeB:A -> TypeR:forall x0 : A, B x0 -> PropH:forall x0 : A, exists ! y0 : B x0, R x0 y0B':={x0 : A & B x0}:TypeR':=fun (x0 : A) (y0 : B') => projT1 y0 = x0 /\ R (projT1 y0) (projT2 y0):A -> B' -> Propx:Ay:B xHy:R x yHuni:forall x'0 : B x, R x x'0 -> y = x'0x':Ay':B x'Heqx':projT1 (existT (fun x0 : A => B x0) x' y') = xHy':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:FunctionalRelReificationA:TypeB:A -> TypeR:forall x : A, B x -> PropH:forall x : A, exists ! y : B x, R x yB':={x : A & B x}:TypeR':=fun (x : A) (y : B') => projT1 y = x /\ R (projT1 y) (projT2 y):A -> B' -> Propf: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:FunctionalRelReificationA:TypeB:A -> TypeR:forall x0 : A, B x0 -> PropH:forall x0 : A, exists ! y0 : B x0, R x0 y0B':={x0 : A & B x0}:TypeR':=fun (x0 : A) (y0 : B') => projT1 y0 = x0 /\ R (projT1 y0) (projT2 y0):A -> B' -> Propx:Ay:B xHy:R x yHuni:forall x'0 : B x, R x x'0 -> y = x'0x':Ay':B x'Heqx':x' = xHy':R x' y'existT (fun x0 : A => B x0) x y = existT (fun x0 : A => B x0) x' y'AC_fun:FunctionalRelReificationA:TypeB:A -> TypeR:forall x : A, B x -> PropH:forall x : A, exists ! y : B x, R x yB':={x : A & B x}:TypeR':=fun (x : A) (y : B') => projT1 y = x /\ R (projT1 y) (projT2 y):A -> B' -> Propf: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:FunctionalRelReificationA:TypeB:A -> TypeR:forall x : A, B x -> PropH:forall x : A, exists ! y0 : B x, R x y0B':={x : A & B x}:TypeR':=fun (x : A) (y0 : B') => projT1 y0 = x /\ R (projT1 y0) (projT2 y0):A -> B' -> Propx':Ay:B x'Hy:R x' yHuni:forall x'0 : B x', R x' x'0 -> y = x'0y':B x'Hy':R x' y'existT (fun x : A => B x) x' y = existT (fun x : A => B x) x' y'AC_fun:FunctionalRelReificationA:TypeB:A -> TypeR:forall x : A, B x -> PropH:forall x : A, exists ! y : B x, R x yB':={x : A & B x}:TypeR':=fun (x : A) (y : B') => projT1 y = x /\ R (projT1 y) (projT2 y):A -> B' -> Propf: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:FunctionalRelReificationA:TypeB:A -> TypeR:forall x : A, B x -> PropH:forall x : A, exists ! y : B x, R x yB':={x : A & B x}:TypeR':=fun (x : A) (y : B') => projT1 y = x /\ R (projT1 y) (projT2 y):A -> B' -> Propf: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:FunctionalRelReificationA:TypeB:A -> TypeR:forall x : A, B x -> PropH:forall x : A, exists ! y : B x, R x yB':={x : A & B x}:TypeR':=fun (x : A) (y : B') => projT1 y = x /\ R (projT1 y) (projT2 y):A -> B' -> Propf: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:FunctionalRelReificationA:TypeB:A -> TypeR:forall x0 : A, B x0 -> PropH:forall x0 : A, exists ! y : B x0, R x0 yB':={x0 : A & B x0}:TypeR':=fun (x0 : A) (y : B') => projT1 y = x0 /\ R (projT1 y) (projT2 y):A -> B' -> Propf:A -> B'Hf:forall x0 : A, R' x0 (f x0)x:AHeq:projT1 (f x) = xHR: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)))destruct Heq using eq_indd; trivial. Qed.AC_fun:FunctionalRelReificationA:TypeB:A -> TypeR:forall x1 : A, B x1 -> PropH:forall x1 : A, exists ! y : B x1, R x1 yB':={x1 : A & B x1}:TypeR':=fun (x1 : A) (y : B') => projT1 y = x1 /\ R (projT1 y) (projT2 y):A -> B' -> Propf:A -> B'Hf:forall x1 : A, R' x1 (f x1)x, x0:Ab:B x0Heq:x0 = xHR:R x0 bR x (eq_rect x0 (fun x1 : A => B x1) b x Heq)FunctionalRelReification <-> DependentFunctionalRelReificationintuition auto using non_dep_dep_functional_rel_reification, dep_non_dep_functional_rel_reification. Qed. (**********************************************************************)FunctionalRelReification <-> DependentFunctionalRelReification
forall C : Prop, (ConstructiveIndefiniteDescription -> C) -> FunctionalChoice -> Cforall C : Prop, (ConstructiveIndefiniteDescription -> C) -> FunctionalChoice -> CC:PropH:ConstructiveIndefiniteDescription -> CAC_fun:FunctionalChoiceCC:PropH:ConstructiveIndefiniteDescription -> CAC_fun:FunctionalChoiceAC_depfun:DependentFunctionalChoiceCC:PropH:ConstructiveIndefiniteDescription -> CAC_fun:FunctionalChoiceAC_depfun:DependentFunctionalChoiceA0:={A : Type & {P : A -> Prop & exists x : A, P x}}:TypeCC:PropH:ConstructiveIndefiniteDescription -> CAC_fun:FunctionalChoiceAC_depfun:DependentFunctionalChoiceA0:={A : Type & {P : A -> Prop & exists x : A, P x}}:TypeB0:=fun x : A0 => projT1 x:A0 -> TypeCC:PropH:ConstructiveIndefiniteDescription -> CAC_fun:FunctionalChoiceAC_depfun:DependentFunctionalChoiceA0:={A : Type & {P : A -> Prop & exists x : A, P x}}:TypeB0:=fun x : A0 => projT1 x:A0 -> TypeR0:=fun (x : A0) (y : B0 x) => projT1 (projT2 x) y:forall x : A0, B0 x -> PropCC:PropH:ConstructiveIndefiniteDescription -> CAC_fun:FunctionalChoiceAC_depfun:DependentFunctionalChoiceA0:={A : Type & {P : A -> Prop & exists x : A, P x}}:TypeB0:=fun x : A0 => projT1 x:A0 -> TypeR0:=fun (x : A0) (y : B0 x) => projT1 (projT2 x) y:forall x : A0, B0 x -> PropH0:=fun x : A0 => projT2 (projT2 x):forall x : A0, (fun P : projT1 x -> Prop => exists x0 : projT1 x, P x0) (projT1 (projT2 x))CC:PropH:ConstructiveIndefiniteDescription -> CAC_fun:FunctionalChoiceAC_depfun:DependentFunctionalChoiceA0:={A : Type & {P : A -> Prop & exists x : A, P x}}:TypeB0:=fun x : A0 => projT1 x:A0 -> TypeR0:=fun (x : A0) (y : B0 x) => projT1 (projT2 x) y:forall x : A0, B0 x -> PropH0:=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 xHf:forall x : A0, R0 x (f x)CC:PropH:ConstructiveIndefiniteDescription -> CAC_fun:FunctionalChoiceAC_depfun:DependentFunctionalChoiceA0:={A : Type & {P : A -> Prop & exists x : A, P x}}:TypeB0:=fun x : A0 => projT1 x:A0 -> TypeR0:=fun (x : A0) (y : B0 x) => projT1 (projT2 x) y:forall x : A0, B0 x -> PropH0:=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 xHf:forall x : A0, R0 x (f x)ConstructiveIndefiniteDescriptionC:PropH:ConstructiveIndefiniteDescription -> CAC_fun:FunctionalChoiceAC_depfun:DependentFunctionalChoiceA0:={A1 : Type & {P0 : A1 -> Prop & exists x : A1, P0 x}}:TypeB0:=fun x : A0 => projT1 x:A0 -> TypeR0:=fun (x : A0) (y : B0 x) => projT1 (projT2 x) y:forall x : A0, B0 x -> PropH0:=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 xHf:forall x : A0, R0 x (f x)A:TypeP:A -> PropH':exists x : A, P x{x : A | P x}C:PropH:ConstructiveIndefiniteDescription -> CAC_fun:FunctionalChoiceAC_depfun:DependentFunctionalChoiceA0:={A1 : Type & {P0 : A1 -> Prop & exists x : A1, P0 x}}:TypeB0:=fun x : A0 => projT1 x:A0 -> TypeR0:=fun (x : A0) (y : B0 x) => projT1 (projT2 x) y:forall x : A0, B0 x -> PropH0:=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 xHf:forall x : A0, R0 x (f x)A:TypeP:A -> PropH':exists x : A, P xP (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.C:PropH:ConstructiveIndefiniteDescription -> CAC_fun:FunctionalChoiceAC_depfun:DependentFunctionalChoiceA0:={A1 : Type & {P0 : A1 -> Prop & exists x : A1, P0 x}}:TypeB0:=fun x : A0 => projT1 x:A0 -> TypeR0:=fun (x : A0) (y : B0 x) => projT1 (projT2 x) y:forall x : A0, B0 x -> PropH0:=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 xHf:forall x : A0, R0 x (f x)A:TypeP:A -> PropH':exists x : A, P xHf':=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')))ConstructiveIndefiniteDescription -> FunctionalChoiceConstructiveIndefiniteDescription -> FunctionalChoiceIndefDescr:ConstructiveIndefiniteDescriptionA, B:TypeR:A -> B -> PropH:forall x : A, exists y : B, R x yexists f : A -> B, forall x : A, R x (f x)IndefDescr:ConstructiveIndefiniteDescriptionA, B:TypeR:A -> B -> PropH:forall x : A, exists y : B, R x yforall x : A, R x (proj1_sig (IndefDescr B (R x) (H x)))apply (proj2_sig (IndefDescr B (R x) (H x))). Qed.IndefDescr:ConstructiveIndefiniteDescriptionA, B:TypeR:A -> B -> PropH:forall x0 : A, exists y : B, R x0 yx:AR x (proj1_sig (IndefDescr B (R x) (H x)))
forall C : Prop, (ConstructiveDefiniteDescription -> C) -> FunctionalRelReification -> Cforall C : Prop, (ConstructiveDefiniteDescription -> C) -> FunctionalRelReification -> CC:PropH:ConstructiveDefiniteDescription -> CFunReify:FunctionalRelReificationCC:PropH:ConstructiveDefiniteDescription -> CFunReify:FunctionalRelReificationDepFunReify:DependentFunctionalRelReificationCC:PropH:ConstructiveDefiniteDescription -> CFunReify:FunctionalRelReificationDepFunReify:DependentFunctionalRelReificationA0:={A : Type & {P : A -> Prop & exists ! x : A, P x}}:TypeCC:PropH:ConstructiveDefiniteDescription -> CFunReify:FunctionalRelReificationDepFunReify:DependentFunctionalRelReificationA0:={A : Type & {P : A -> Prop & exists ! x : A, P x}}:TypeB0:=fun x : A0 => projT1 x:A0 -> TypeCC:PropH:ConstructiveDefiniteDescription -> CFunReify:FunctionalRelReificationDepFunReify:DependentFunctionalRelReificationA0:={A : Type & {P : A -> Prop & exists ! x : A, P x}}:TypeB0:=fun x : A0 => projT1 x:A0 -> TypeR0:=fun (x : A0) (y : B0 x) => projT1 (projT2 x) y:forall x : A0, B0 x -> PropCC:PropH:ConstructiveDefiniteDescription -> CFunReify:FunctionalRelReificationDepFunReify:DependentFunctionalRelReificationA0:={A : Type & {P : A -> Prop & exists ! x : A, P x}}:TypeB0:=fun x : A0 => projT1 x:A0 -> TypeR0:=fun (x : A0) (y : B0 x) => projT1 (projT2 x) y:forall x : A0, B0 x -> PropH0:=fun x : A0 => projT2 (projT2 x):forall x : A0, (fun P : projT1 x -> Prop => exists ! x0 : projT1 x, P x0) (projT1 (projT2 x))CC:PropH:ConstructiveDefiniteDescription -> CFunReify:FunctionalRelReificationDepFunReify:DependentFunctionalRelReificationA0:={A : Type & {P : A -> Prop & exists ! x : A, P x}}:TypeB0:=fun x : A0 => projT1 x:A0 -> TypeR0:=fun (x : A0) (y : B0 x) => projT1 (projT2 x) y:forall x : A0, B0 x -> PropH0:=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 xHf:forall x : A0, R0 x (f x)CC:PropH:ConstructiveDefiniteDescription -> CFunReify:FunctionalRelReificationDepFunReify:DependentFunctionalRelReificationA0:={A : Type & {P : A -> Prop & exists ! x : A, P x}}:TypeB0:=fun x : A0 => projT1 x:A0 -> TypeR0:=fun (x : A0) (y : B0 x) => projT1 (projT2 x) y:forall x : A0, B0 x -> PropH0:=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 xHf:forall x : A0, R0 x (f x)ConstructiveDefiniteDescriptionC:PropH:ConstructiveDefiniteDescription -> CFunReify:FunctionalRelReificationDepFunReify:DependentFunctionalRelReificationA0:={A1 : Type & {P0 : A1 -> Prop & exists ! x : A1, P0 x}}:TypeB0:=fun x : A0 => projT1 x:A0 -> TypeR0:=fun (x : A0) (y : B0 x) => projT1 (projT2 x) y:forall x : A0, B0 x -> PropH0:=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 xHf:forall x : A0, R0 x (f x)A:TypeP:A -> PropH':exists ! x : A, P x{x : A | P x}C:PropH:ConstructiveDefiniteDescription -> CFunReify:FunctionalRelReificationDepFunReify:DependentFunctionalRelReificationA0:={A1 : Type & {P0 : A1 -> Prop & exists ! x : A1, P0 x}}:TypeB0:=fun x : A0 => projT1 x:A0 -> TypeR0:=fun (x : A0) (y : B0 x) => projT1 (projT2 x) y:forall x : A0, B0 x -> PropH0:=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 xHf:forall x : A0, R0 x (f x)A:TypeP:A -> PropH':exists ! x : A, P xP (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.C:PropH:ConstructiveDefiniteDescription -> CFunReify:FunctionalRelReificationDepFunReify:DependentFunctionalRelReificationA0:={A1 : Type & {P0 : A1 -> Prop & exists ! x : A1, P0 x}}:TypeB0:=fun x : A0 => projT1 x:A0 -> TypeR0:=fun (x : A0) (y : B0 x) => projT1 (projT2 x) y:forall x : A0, B0 x -> PropH0:=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 xHf:forall x : A0, R0 x (f x)A:TypeP:A -> PropH':exists ! x : A, P xHf':=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')))ConstructiveDefiniteDescription -> FunctionalRelReificationConstructiveDefiniteDescription -> FunctionalRelReificationDefDescr:ConstructiveDefiniteDescriptionA, B:TypeR:A -> B -> PropH:forall x : A, exists ! y : B, R x yexists f : A -> B, forall x : A, R x (f x)DefDescr:ConstructiveDefiniteDescriptionA, B:TypeR:A -> B -> PropH:forall x : A, exists ! y : B, R x yforall x : A, R x (proj1_sig (DefDescr B (R x) (H x)))apply (proj2_sig (DefDescr B (R x) (H x))). Qed.DefDescr:ConstructiveDefiniteDescriptionA, B:TypeR:A -> B -> PropH:forall x0 : A, exists ! y : B, R x0 yx:AR x (proj1_sig (DefDescr B (R x) (H x)))
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.
(**********************************************************************)
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:ConstructiveDefiniteDescriptionEM:forall P0 : Prop, P0 \/ ~ P0P:Prop{P} + {~ P}Descr:ConstructiveDefiniteDescriptionEM:forall P0 : Prop, P0 \/ ~ P0P:Propselect:=fun b : bool => if b then P else ~ P:bool -> Prop{P} + {~ P}Descr:ConstructiveDefiniteDescriptionEM:forall P0 : Prop, P0 \/ ~ P0P:Propselect:=fun b : bool => if b then P else ~ P:bool -> Prop{b : bool | select b}Descr:ConstructiveDefiniteDescriptionEM:forall P0 : Prop, P0 \/ ~ P0P:Propselect:=fun b : bool => if b then P else ~ P:bool -> PropHP:select true{P} + {~ P}Descr:ConstructiveDefiniteDescriptionEM:forall P0 : Prop, P0 \/ ~ P0P:Propselect:=fun b : bool => if b then P else ~ P:bool -> PropHP: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 \/ ~ P0P:Propselect:=fun b : bool => if b then P else ~ P:bool -> Prop{b : bool | select b}Descr:ConstructiveDefiniteDescriptionEM:forall P0 : Prop, P0 \/ ~ P0P:Propselect:=fun b : bool => if b then P else ~ P:bool -> PropHP:select true{P} + {~ P}Descr:ConstructiveDefiniteDescriptionEM:forall P0 : Prop, P0 \/ ~ P0P:Propselect:=fun b : bool => if b then P else ~ P:bool -> PropHP: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 \/ ~ P0P:Propselect:=fun b : bool => if b then P else ~ P:bool -> Propexists ! x : bool, select xDescr:ConstructiveDefiniteDescriptionEM:forall P0 : Prop, P0 \/ ~ P0P:Propselect:=fun b : bool => if b then P else ~ P:bool -> PropHP:select true{P} + {~ P}Descr:ConstructiveDefiniteDescriptionEM:forall P0 : Prop, P0 \/ ~ P0P:Propselect:=fun b : bool => if b then P else ~ P:bool -> PropHP: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 \/ ~ P0P:Propselect:=fun b : bool => if b then P else ~ P:bool -> Propexists x : bool, select xDescr:forall (A : Type) (P0 : A -> Prop), (exists ! x : A, P0 x) -> {x : A | P0 x}EM:forall P0 : Prop, P0 \/ ~ P0P:Propselect:=fun b : bool => if b then P else ~ P:bool -> Propuniqueness selectDescr:ConstructiveDefiniteDescriptionEM:forall P0 : Prop, P0 \/ ~ P0P:Propselect:=fun b : bool => if b then P else ~ P:bool -> PropHP:select true{P} + {~ P}Descr:ConstructiveDefiniteDescriptionEM:forall P0 : Prop, P0 \/ ~ P0P:Propselect:=fun b : bool => if b then P else ~ P:bool -> PropHP: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 \/ ~ P0P:Propselect:=fun b : bool => if b then P else ~ P:bool -> PropH:Pexists x : bool, select xDescr:forall (A : Type) (P0 : A -> Prop), (exists ! x : A, P0 x) -> {x : A | P0 x}EM:forall P0 : Prop, P0 \/ ~ P0P:Propselect:=fun b : bool => if b then P else ~ P:bool -> PropH:~ Pexists x : bool, select xDescr:forall (A : Type) (P0 : A -> Prop), (exists ! x : A, P0 x) -> {x : A | P0 x}EM:forall P0 : Prop, P0 \/ ~ P0P:Propselect:=fun b : bool => if b then P else ~ P:bool -> Propuniqueness selectDescr:ConstructiveDefiniteDescriptionEM:forall P0 : Prop, P0 \/ ~ P0P:Propselect:=fun b : bool => if b then P else ~ P:bool -> PropHP:select true{P} + {~ P}Descr:ConstructiveDefiniteDescriptionEM:forall P0 : Prop, P0 \/ ~ P0P:Propselect:=fun b : bool => if b then P else ~ P:bool -> PropHP: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 \/ ~ P0P:Propselect:=fun b : bool => if b then P else ~ P:bool -> PropH:~ Pexists x : bool, select xDescr:forall (A : Type) (P0 : A -> Prop), (exists ! x : A, P0 x) -> {x : A | P0 x}EM:forall P0 : Prop, P0 \/ ~ P0P:Propselect:=fun b : bool => if b then P else ~ P:bool -> Propuniqueness selectDescr:ConstructiveDefiniteDescriptionEM:forall P0 : Prop, P0 \/ ~ P0P:Propselect:=fun b : bool => if b then P else ~ P:bool -> PropHP:select true{P} + {~ P}Descr:ConstructiveDefiniteDescriptionEM:forall P0 : Prop, P0 \/ ~ P0P:Propselect:=fun b : bool => if b then P else ~ P:bool -> PropHP: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 \/ ~ P0P:Propselect:=fun b : bool => if b then P else ~ P:bool -> Propuniqueness selectDescr:ConstructiveDefiniteDescriptionEM:forall P0 : Prop, P0 \/ ~ P0P:Propselect:=fun b : bool => if b then P else ~ P:bool -> PropHP:select true{P} + {~ P}Descr:ConstructiveDefiniteDescriptionEM:forall P0 : Prop, P0 \/ ~ P0P:Propselect:=fun b : bool => if b then P else ~ P:bool -> PropHP:select false{P} + {~ P}Descr:ConstructiveDefiniteDescriptionEM:forall P0 : Prop, P0 \/ ~ P0P:Propselect:=fun b : bool => if b then P else ~ P:bool -> PropHP:select true{P} + {~ P}Descr:ConstructiveDefiniteDescriptionEM:forall P0 : Prop, P0 \/ ~ P0P:Propselect:=fun b : bool => if b then P else ~ P:bool -> PropHP:select false{P} + {~ P}right; trivial. Qed.Descr:ConstructiveDefiniteDescriptionEM:forall P0 : Prop, P0 \/ ~ P0P:Propselect:=fun b : bool => if b then P else ~ P:bool -> PropHP:select false{P} + {~ P}FunctionalRelReification -> (forall P : Prop, P \/ ~ P) -> forall C : Prop, ((forall P : Prop, {P} + {~ P}) -> C) -> CFunctionalRelReification -> (forall P : Prop, P \/ ~ P) -> forall C : Prop, ((forall P : Prop, {P} + {~ P}) -> C) -> Cintuition auto using constructive_definite_descr_excluded_middle, (relative_non_contradiction_of_definite_descr (C:=C)). Qed. (**********************************************************************)FunReify:FunctionalRelReificationEM:forall P : Prop, P \/ ~ PC:PropH:(forall P : Prop, {P} + {~ P}) -> CC
(* The implications below are standard *) Require Import Arith.FunctionalChoice -> FunctionalDependentChoiceFunctionalChoice -> FunctionalDependentChoiceFunChoice:FunctionalChoiceA:TypeR:A -> A -> PropHRfun:forall x : A, exists y : A, R x yx0:Aexists f : nat -> A, f 0 = x0 /\ (forall n : nat, R (f n) (f (S n)))FunChoice:FunctionalChoiceA:TypeR:A -> A -> Propg:A -> ARg:forall x : A, R x (g x)x0:Aexists f : nat -> A, f 0 = x0 /\ (forall n : nat, R (f n) (f (S n)))exists f; firstorder. Qed.FunChoice:FunctionalChoiceA:TypeR:A -> A -> Propg:A -> ARg:forall x : A, R x (g x)x0:Af:=fix f0 (n : nat) : A := match n with | 0 => x0 | S n' => g (f0 n') end:nat -> Aexists f0 : nat -> A, f0 0 = x0 /\ (forall n : nat, R (f0 n) (f0 (S n)))FunctionalDependentChoice -> FunctionalCountableChoiceFunctionalDependentChoice -> FunctionalCountableChoiceH:FunctionalDependentChoiceA:TypeR:nat -> A -> PropH0:forall n : nat, exists y : A, R n yexists f : nat -> A, forall n : nat, R n (f n)H:FunctionalDependentChoiceA:TypeR:nat -> A -> PropH0:forall n : nat, exists y : A, R n yR':=fun p q : nat * A => fst q = S (fst p) /\ R (fst p) (snd q):nat * A -> nat * A -> Propexists f : nat -> A, forall n : nat, R n (f n)H:FunctionalDependentChoiceA:TypeR:nat -> A -> PropH0:forall n : nat, exists y : A, R n yR':=fun p q : nat * A => fst q = S (fst p) /\ R (fst p) (snd q):nat * A -> nat * A -> Propy0:AHy0:R 0 y0exists f : nat -> A, forall n : nat, R n (f n)H:FunctionalDependentChoiceA:TypeR:nat -> A -> PropH0:forall n : nat, exists y : A, R n yR':=fun p q : nat * A => fst q = S (fst p) /\ R (fst p) (snd q):nat * A -> nat * A -> Propy0:AHy0:R 0 y0forall x : nat * A, exists y : nat * A, R' x yH:FunctionalDependentChoiceA:TypeR:nat -> A -> PropH0:forall n : nat, exists y : A, R n yR':=fun p q : nat * A => fst q = S (fst p) /\ R (fst p) (snd q):nat * A -> nat * A -> Propy0:AHy0:R 0 y0f:nat -> nat * AHf0: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:FunctionalDependentChoiceA:TypeR:nat -> A -> PropH0:forall n : nat, exists y1 : A, R n y1R':=fun p q : nat * A => fst q = S (fst p) /\ R (fst p) (snd q):nat * A -> nat * A -> Propy0:AHy0:R 0 y0x:(nat * A)%typey:AHy:R (fst x) yexists y1 : nat * A, R' x y1H:FunctionalDependentChoiceA:TypeR:nat -> A -> PropH0:forall n : nat, exists y : A, R n yR':=fun p q : nat * A => fst q = S (fst p) /\ R (fst p) (snd q):nat * A -> nat * A -> Propy0:AHy0:R 0 y0f:nat -> nat * AHf0: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:FunctionalDependentChoiceA:TypeR:nat -> A -> PropH0:forall n : nat, exists y1 : A, R n y1R':=fun p q : nat * A => fst q = S (fst p) /\ R (fst p) (snd q):nat * A -> nat * A -> Propy0:AHy0:R 0 y0x:(nat * A)%typey:AHy:R (fst x) yR' x (S (fst x), y)H:FunctionalDependentChoiceA:TypeR:nat -> A -> PropH0:forall n : nat, exists y : A, R n yR':=fun p q : nat * A => fst q = S (fst p) /\ R (fst p) (snd q):nat * A -> nat * A -> Propy0:AHy0:R 0 y0f:nat -> nat * AHf0: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:FunctionalDependentChoiceA:TypeR:nat -> A -> PropH0:forall n : nat, exists y1 : A, R n y1R':=fun p q : nat * A => fst q = S (fst p) /\ R (fst p) (snd q):nat * A -> nat * A -> Propy0:AHy0:R 0 y0x:(nat * A)%typey:AHy:R (fst x) yfst (S (fst x), y) = S (fst x) /\ R (fst x) (snd (S (fst x), y))H:FunctionalDependentChoiceA:TypeR:nat -> A -> PropH0:forall n : nat, exists y : A, R n yR':=fun p q : nat * A => fst q = S (fst p) /\ R (fst p) (snd q):nat * A -> nat * A -> Propy0:AHy0:R 0 y0f:nat -> nat * AHf0: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:FunctionalDependentChoiceA:TypeR:nat -> A -> PropH0:forall n : nat, exists y : A, R n yR':=fun p q : nat * A => fst q = S (fst p) /\ R (fst p) (snd q):nat * A -> nat * A -> Propy0:AHy0:R 0 y0f:nat -> nat * AHf0: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:FunctionalDependentChoiceA:TypeR:nat -> A -> PropH0:forall n : nat, exists y : A, R n yR':=fun p q : nat * A => fst q = S (fst p) /\ R (fst p) (snd q):nat * A -> nat * A -> Propy0:AHy0:R 0 y0f:nat -> nat * AHf0:f 0 = (0, y0)HfS:forall n : nat, R' (f n) (f (S n))forall n : nat, fst (f n) = nH:FunctionalDependentChoiceA:TypeR:nat -> A -> PropH0:forall n : nat, exists y : A, R n yR':=fun p q : nat * A => fst q = S (fst p) /\ R (fst p) (snd q):nat * A -> nat * A -> Propy0:AHy0:R 0 y0f:nat -> nat * AHf0:f 0 = (0, y0)HfS:forall n : nat, R' (f n) (f (S n))Heq:forall n : nat, fst (f n) = nexists f0 : nat -> A, forall n : nat, R n (f0 n)H:FunctionalDependentChoiceA:TypeR:nat -> A -> PropH0:forall n : nat, exists y : A, R n yR':=fun p q : nat * A => fst q = S (fst p) /\ R (fst p) (snd q):nat * A -> nat * A -> Propy0:AHy0:R 0 y0f:nat -> nat * AHf0:f 0 = (0, y0)HfS:forall n : nat, R' (f n) (f (S n))fst (f 0) = 0H:FunctionalDependentChoiceA:TypeR:nat -> A -> PropH0:forall n0 : nat, exists y : A, R n0 yR':=fun p q : nat * A => fst q = S (fst p) /\ R (fst p) (snd q):nat * A -> nat * A -> Propy0:AHy0:R 0 y0f:nat -> nat * AHf0:f 0 = (0, y0)HfS:forall n0 : nat, R' (f n0) (f (S n0))n:natIHn:fst (f n) = nfst (f (S n)) = S nH:FunctionalDependentChoiceA:TypeR:nat -> A -> PropH0:forall n : nat, exists y : A, R n yR':=fun p q : nat * A => fst q = S (fst p) /\ R (fst p) (snd q):nat * A -> nat * A -> Propy0:AHy0:R 0 y0f:nat -> nat * AHf0:f 0 = (0, y0)HfS:forall n : nat, R' (f n) (f (S n))Heq:forall n : nat, fst (f n) = nexists f0 : nat -> A, forall n : nat, R n (f0 n)H:FunctionalDependentChoiceA:TypeR:nat -> A -> PropH0:forall n0 : nat, exists y : A, R n0 yR':=fun p q : nat * A => fst q = S (fst p) /\ R (fst p) (snd q):nat * A -> nat * A -> Propy0:AHy0:R 0 y0f:nat -> nat * AHf0:f 0 = (0, y0)HfS:forall n0 : nat, R' (f n0) (f (S n0))n:natIHn:fst (f n) = nfst (f (S n)) = S nH:FunctionalDependentChoiceA:TypeR:nat -> A -> PropH0:forall n : nat, exists y : A, R n yR':=fun p q : nat * A => fst q = S (fst p) /\ R (fst p) (snd q):nat * A -> nat * A -> Propy0:AHy0:R 0 y0f:nat -> nat * AHf0:f 0 = (0, y0)HfS:forall n : nat, R' (f n) (f (S n))Heq:forall n : nat, fst (f n) = nexists f0 : nat -> A, forall n : nat, R n (f0 n)H:FunctionalDependentChoiceA:TypeR:nat -> A -> PropH0:forall n : nat, exists y : A, R n yR':=fun p q : nat * A => fst q = S (fst p) /\ R (fst p) (snd q):nat * A -> nat * A -> Propy0:AHy0:R 0 y0f:nat -> nat * AHf0:f 0 = (0, y0)HfS:forall n : nat, R' (f n) (f (S n))Heq:forall n : nat, fst (f n) = nexists f0 : nat -> A, forall n : nat, R n (f0 n)H:FunctionalDependentChoiceA:TypeR:nat -> A -> PropH0:forall n : nat, exists y : A, R n yR':=fun p q : nat * A => fst q = S (fst p) /\ R (fst p) (snd q):nat * A -> nat * A -> Propy0:AHy0:R 0 y0f:nat -> nat * AHf0:f 0 = (0, y0)HfS:forall n : nat, R' (f n) (f (S n))Heq:forall n : nat, fst (f n) = nforall n : nat, R n (snd (f (S n)))H:FunctionalDependentChoiceA:TypeR:nat -> A -> PropH0:forall n : nat, exists y : A, R n yR':=fun p q : nat * A => fst q = S (fst p) /\ R (fst p) (snd q):nat * A -> nat * A -> Propy0:AHy0:R 0 y0f:nat -> nat * AHf0:f 0 = (0, y0)HfS:forall n : nat, R' (f n) (f (S n))Heq:forall n : nat, fst (f n) = nn':natR n' (snd (f (S n')))H:FunctionalDependentChoiceA:TypeR:nat -> A -> PropH0:forall n : nat, exists y : A, R n yR':=fun p q : nat * A => fst q = S (fst p) /\ R (fst p) (snd q):nat * A -> nat * A -> Propy0:AHy0:R 0 y0f:nat -> nat * AHf0:f 0 = (0, y0)n':natHfS:R' (f n') (f (S n'))Heq:forall n : nat, fst (f n) = nR n' (snd (f (S n')))H:FunctionalDependentChoiceA:TypeR:nat -> A -> PropH0:forall n : nat, exists y : A, R n yR':=fun p q : nat * A => fst q = S (fst p) /\ R (fst p) (snd q):nat * A -> nat * A -> Propy0:AHy0:R 0 y0f:nat -> nat * AHf0:f 0 = (0, y0)n':natHR:R (fst (f n')) (snd (f (S n')))Heq:forall n : nat, fst (f n) = nR n' (snd (f (S n')))assumption. Qed. (**********************************************************************)H:FunctionalDependentChoiceA:TypeR:nat -> A -> PropH0:forall n : nat, exists y : A, R n yR':=fun p q : nat * A => fst q = S (fst p) /\ R (fst p) (snd q):nat * A -> nat * A -> Propy0:AHy0:R 0 y0f:nat -> nat * AHf0:f 0 = (0, y0)n':natHR:R n' (snd (f (S n')))Heq:forall n : nat, fst (f n) = nR n' (snd (f (S n')))
Require Import ClassicalFacts PropExtensionalityFacts. (**********************************************************************)
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:RepresentativeFunctionalChoiceA:Typeexists h : Prop -> Prop, forall P : Prop, (P <-> h P) /\ (forall Q : Prop, P <-> Q -> h P = h Q)ReprFunChoice:RepresentativeFunctionalChoiceA:TypeR:=fun P Q : Prop => P <-> Q:Prop -> Prop -> Propexists h : Prop -> Prop, forall P : Prop, (P <-> h P) /\ (forall Q : Prop, P <-> Q -> h P = h Q)apply (ReprFunChoice _ R Hequiv). Qed.ReprFunChoice:RepresentativeFunctionalChoiceA:TypeR:=fun P Q : Prop => P <-> Q:Prop -> Prop -> PropHequiv:Equivalence Rexists h : Prop -> Prop, forall P : Prop, (P <-> h P) /\ (forall Q : Prop, P <-> Q -> h P = h Q)RepresentativeFunctionalChoice -> ExtensionalPredicateRepresentativeRepresentativeFunctionalChoice -> ExtensionalPredicateRepresentativeReprFunChoice:RepresentativeFunctionalChoiceA:Typeexists 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:RepresentativeFunctionalChoiceA:TypeR:=fun P Q : A -> Prop => forall x : A, P x <-> Q x:(A -> Prop) -> (A -> Prop) -> Propexists 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.ReprFunChoice:RepresentativeFunctionalChoiceA:TypeR:=fun P Q : A -> Prop => forall x : A, P x <-> Q x:(A -> Prop) -> (A -> Prop) -> PropHequiv:Equivalence Rexists 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)RepresentativeFunctionalChoice -> ExtensionalFunctionRepresentativeRepresentativeFunctionalChoice -> ExtensionalFunctionRepresentativeReprFunChoice:RepresentativeFunctionalChoiceA, B:Typeexists 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:RepresentativeFunctionalChoiceA, B:TypeR:=fun f g : A -> B => forall x : A, f x = g x:(A -> B) -> (A -> B) -> Propexists 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:RepresentativeFunctionalChoiceA, B:TypeR:=fun f g : A -> B => forall x : A, f x = g x:(A -> B) -> (A -> B) -> PropEquivalence RReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeR:=fun f g : A -> B => forall x : A, f x = g x:(A -> B) -> (A -> B) -> PropHequiv:Equivalence Rexists 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:RepresentativeFunctionalChoiceA, B:TypeR:=fun f g : A -> B => forall x : A, f x = g x:(A -> B) -> (A -> B) -> PropEquivalence Rfirstorder using eq_trans.ReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeR:=fun f g : A -> B => forall x : A, f x = g x:(A -> B) -> (A -> B) -> PropTransitive Rapply (ReprFunChoice _ R Hequiv). Qed.ReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeR:=fun f g : A -> B => forall x : A, f x = g x:(A -> B) -> (A -> B) -> PropHequiv:Equivalence Rexists 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)
RepresentativeFunctionalChoice -> ExcludedMiddleRepresentativeFunctionalChoice -> ExcludedMiddleapply representative_boolean_partition_imp_excluded_middle, ReprFunChoice. Qed.ReprFunChoice:RepresentativeFunctionalChoiceExcludedMiddleRepresentativeFunctionalChoice -> RelationalChoiceRepresentativeFunctionalChoice -> RelationalChoiceReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x : A, exists y : B, T x yexists R' : A -> B -> Prop, subrelation R' T /\ (forall x : A, exists ! y : B, R' x y)ReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x : A, exists y : B, T x yD:=(A * B)%type:Typeexists R' : A -> B -> Prop, subrelation R' T /\ (forall x : A, exists ! y : B, R' x y)ReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x : A, exists y : B, T x yD:=(A * B)%type:TypeR:=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 -> Propexists R' : A -> B -> Prop, subrelation R' T /\ (forall x : A, exists ! y : B, R' x y)ReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x : A, exists y : B, T x yD:=(A * B)%type:TypeR:=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 -> PropEquivalence RReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x : A, exists y : B, T x yD:=(A * B)%type:TypeR:=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 -> PropHequiv:Equivalence Rexists R' : A -> B -> Prop, subrelation R' T /\ (forall x : A, exists ! y : B, R' x y)ReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x : A, exists y : B, T x yD:=(A * B)%type:TypeR:=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 -> PropEquivalence RReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x : A, exists y : B, T x yD:=(A * B)%type:TypeR:=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 -> PropReflexive RReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x : A, exists y : B, T x yD:=(A * B)%type:TypeR:=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 -> PropSymmetric RReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x : A, exists y : B, T x yD:=(A * B)%type:TypeR:=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 -> PropTransitive RReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x : A, exists y : B, T x yD:=(A * B)%type:TypeR:=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 -> PropReflexive RReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y : B, T x0 yD:=(A * B)%type:TypeR:=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 -> Propx:Dfst x = fst xReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y : B, T x0 yD:=(A * B)%type:TypeR:=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 -> Propx: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:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y : B, T x0 yD:=(A * B)%type:TypeR:=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 -> Propx: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:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x : A, exists y : B, T x yD:=(A * B)%type:TypeR:=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 -> PropSymmetric RReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> Propx:Ay:Bx':Ay':BH1: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:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> Propx:Ay:Bx':Ay':BH1: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:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> Propx:Ay:Bx':Ay':BH1: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:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> Propx:Ay:Bx':Ay':BH1: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:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> Propx:Ay:Bx':Ay':BH1: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:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> Propx:Ay:Bx':Ay':BH1: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:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> Propx:Ay, y':BH2: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:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> Propx:Ay, y':BH2:T x y -> y = y' \/ T x y'H2':T x y' -> y = y' \/ T x yH:T x y'y' = y \/ T x yReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> Propx:Ay, y':BH2:T x y -> y = y' \/ T x y'H2':T x y' -> y = y' \/ T x yH:T x yy' = y \/ T x y'destruct (H2' H); firstorder.ReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> Propx:Ay, y':BH2:T x y -> y = y' \/ T x y'H2':T x y' -> y = y' \/ T x yH:T x y'y' = y \/ T x ydestruct (H2 H); firstorder.ReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> Propx:Ay, y':BH2:T x y -> y = y' \/ T x y'H2':T x y' -> y = y' \/ T x yH:T x yy' = y \/ T x y'ReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x : A, exists y : B, T x yD:=(A * B)%type:TypeR:=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 -> PropTransitive RReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> Propx:Ay:Bx':Ay':Bx'':Ay'':BH1: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:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> Propx:Ay:Bx':Ay':Bx'':Ay'':BH1: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:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> Propx:Ay:Bx':Ay':Bx'':Ay'':BH1:x = x'H2:T x y -> y = y' \/ T x y'H2':T x y' -> y = y' \/ T x yH3: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:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> Propx:Ay, y', y'':BH2:T x y -> y = y' \/ T x y'H2':T x y' -> y = y' \/ T x yH4':T x y'' -> y' = y'' \/ T x y'H4:T x y' -> y' = y'' \/ T x y''R (x, y) (x, y'')ReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> Propx:Ay, y', y'':BH2:T x y -> y = y' \/ T x y'H2':T x y' -> y = y' \/ T x yH4':T x y'' -> y' = y'' \/ T x y'H4:T x y' -> y' = y'' \/ T x y''fst (x, y) = fst (x, y'')ReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> Propx:Ay, y', y'':BH2:T x y -> y = y' \/ T x y'H2':T x y' -> y = y' \/ T x yH4':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:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> Propx:Ay, y', y'':BH2:T x y -> y = y' \/ T x y'H2':T x y' -> y = y' \/ T x yH4':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:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> Propx:Ay, y', y'':BH2:T x y -> y = y' \/ T x y'H2':T x y' -> y = y' \/ T x yH4':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:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> Propx:Ay, y', y'':BH2:T x y -> y = y' \/ T x y'H2':T x y' -> y = y' \/ T x yH4':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:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> Propx:Ay, y', y'':BH2:T x y -> y = y' \/ T x y'H2':T x y' -> y = y' \/ T x yH4':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:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> Propx:Ay, y', y'':BH2:T x y -> y = y' \/ T x y'H2':T x y' -> y = y' \/ T x yH4':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:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> Propx:Ay, y', y'':BH2:T x y -> y = y' \/ T x y'H2':T x y' -> y = y' \/ T x yH4':T x y'' -> y' = y'' \/ T x y'H4:T x y' -> y' = y'' \/ T x y''H:T x yy = y'' \/ T x y''ReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> Propx:Ay, y'':BH4:T x y -> y = y'' \/ T x y''H4':T x y'' -> y = y'' \/ T x yH2', H2:T x y -> y = y \/ T x yH:T x yy = y'' \/ T x y''ReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> Propx:Ay, y', y'':BH2:T x y -> y = y' \/ T x y'H2':T x y' -> y = y' \/ T x yH4':T x y'' -> y' = y'' \/ T x y'H4:T x y' -> y' = y'' \/ T x y''H:T x yH0:T x y'y = y'' \/ T x y''destruct (H4 H); firstorder.ReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> Propx:Ay, y'':BH4:T x y -> y = y'' \/ T x y''H4':T x y'' -> y = y'' \/ T x yH2', H2:T x y -> y = y \/ T x yH:T x yy = y'' \/ T x y''destruct (H2' H0), (H4 H0); try subst y'; try subst y''; try firstorder.ReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> Propx:Ay, y', y'':BH2:T x y -> y = y' \/ T x y'H2':T x y' -> y = y' \/ T x yH4':T x y'' -> y' = y'' \/ T x y'H4:T x y' -> y' = y'' \/ T x y''H:T x yH0:T x y'y = y'' \/ T x y''ReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> Propx:Ay, y', y'':BH2:T x y -> y = y' \/ T x y'H2':T x y' -> y = y' \/ T x yH4':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:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> Propx:Ay, y', y'':BH2:T x y -> y = y' \/ T x y'H2':T x y' -> y = y' \/ T x yH4':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:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> Propx:Ay, y', y'':BH2:T x y -> y = y' \/ T x y'H2':T x y' -> y = y' \/ T x yH4':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 yReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> Propx:Ay, y':BH2:T x y -> y = y' \/ T x y'H2':T x y' -> y = y' \/ T x yH:T x y'H4, H4':T x y' -> y' = y' \/ T x y'y = y' \/ T x yReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> Propx:Ay, y', y'':BH2:T x y -> y = y' \/ T x y'H2':T x y' -> y = y' \/ T x yH4':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 ydestruct (H2' H); firstorder.ReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> Propx:Ay, y':BH2:T x y -> y = y' \/ T x y'H2':T x y' -> y = y' \/ T x yH:T x y'H4, H4':T x y' -> y' = y' \/ T x y'y = y' \/ T x ydestruct (H2' H0), (H4 H0); try subst y'; try subst y''; try firstorder.ReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> Propx:Ay, y', y'':BH2:T x y -> y = y' \/ T x y'H2':T x y' -> y = y' \/ T x yH4':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 yReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x : A, exists y : B, T x yD:=(A * B)%type:TypeR:=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 -> PropHequiv:Equivalence Rexists R' : A -> B -> Prop, subrelation R' T /\ (forall x : A, exists ! y : B, R' x y)ReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x : A, exists y : B, T x yD:=(A * B)%type:TypeR:=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 -> PropHequiv:Equivalence Rg:D -> DHg: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:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x : A, exists y : B, T x yD:=(A * B)%type:TypeR:=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 -> PropHequiv:Equivalence Rg:D -> DHg: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 -> Propexists R' : A -> B -> Prop, subrelation R' T /\ (forall x : A, exists ! y : B, R' x y)ReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x : A, exists y : B, T x yD:=(A * B)%type:TypeR:=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 -> PropHequiv:Equivalence Rg:D -> DHg: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 -> Propsubrelation T' T /\ (forall x : A, exists ! y : B, T' x y)ReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x : A, exists y : B, T x yD:=(A * B)%type:TypeR:=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 -> PropHequiv:Equivalence Rg:D -> DHg: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 -> Propsubrelation T' TReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x : A, exists y : B, T x yD:=(A * B)%type:TypeR:=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 -> PropHequiv:Equivalence Rg:D -> DHg: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 -> Propforall x : A, exists ! y : B, T' x yintros x y (H,_); easy.ReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x : A, exists y : B, T x yD:=(A * B)%type:TypeR:=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 -> PropHequiv:Equivalence Rg:D -> DHg: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 -> Propsubrelation T' TReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x : A, exists y : B, T x yD:=(A * B)%type:TypeR:=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 -> PropHequiv:Equivalence Rg:D -> DHg: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 -> Propforall x : A, exists ! y : B, T' x yReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y : B, T x0 yD:=(A * B)%type:TypeR:=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 -> PropHequiv:Equivalence Rg:D -> DHg: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 -> Propx:Aexists ! y : B, T' x yReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> PropHequiv:Equivalence Rg:D -> DHg: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 -> Propx:Ay:BHy:T x yexists ! y0 : B, T' x y0ReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> PropHequiv:Equivalence Rg:D -> DHg: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 -> Propx:Ay:BHy:T x yunique (fun y0 : B => T' x y0) (snd (g (x, y)))ReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> PropHequiv:Equivalence Rg:D -> DT':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y' : B, T x0 y' /\ g (x0, y') = (x0, y0)):A -> B -> Propx:Ay:BHy:T x yHeq1: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:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> PropHequiv:Equivalence Rg:D -> DT':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y' : B, T x0 y' /\ g (x0, y') = (x0, y0)):A -> B -> Propx:Ay:BHy:T x yHeq1: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:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> PropHequiv:Equivalence Rg:D -> DT':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y' : B, T x0 y' /\ g (x0, y') = (x0, y0)):A -> B -> Propx:Ay:BHy:T x yHeq1: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:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> PropHequiv:Equivalence Rg:D -> DT':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y' : B, T x0 y' /\ g (x0, y') = (x0, y0)):A -> B -> Propx:Ay:BHy:T x yHeq1: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:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> PropHequiv:Equivalence Rg:D -> DT':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y' : B, T x0 y' /\ g (x0, y') = (x0, y0)):A -> B -> Propx:Ay:BHy:T x yHeq1: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:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> PropHequiv:Equivalence Rg:D -> DT':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y' : B, T x0 y' /\ g (x0, y') = (x0, y0)):A -> B -> Propx:Ay:BHy:T x yHeq1: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:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> PropHequiv:Equivalence Rg:D -> DT':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y' : B, T x0 y' /\ g (x0, y') = (x0, y0)):A -> B -> Propx:Ay:BHy:T x yHeq1: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:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> PropHequiv:Equivalence Rg:D -> DT':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y' : B, T x0 y' /\ g (x0, y') = (x0, y0)):A -> B -> Propx:Ay:BHy:T x yHeq1: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:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> PropHequiv:Equivalence Rg:D -> DT':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y' : B, T x0 y' /\ g (x0, y') = (x0, y0)):A -> B -> Propx:Ay:BHy:T x yHeq1: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:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> PropHequiv:Equivalence Rg:D -> DT':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y' : B, T x0 y' /\ g (x0, y') = (x0, y0)):A -> B -> Propx:Ay:BHy:T x yHeq1: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)))assumption.ReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> PropHequiv:Equivalence Rg:D -> DT':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y' : B, T x0 y' /\ g (x0, y') = (x0, y0)):A -> B -> Propx:Ay:BHy:T x yHeq1: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))ReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> PropHequiv:Equivalence Rg:D -> DT':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y' : B, T x0 y' /\ g (x0, y') = (x0, y0)):A -> B -> Propx:Ay:BHy:T x yHeq1: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:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> PropHequiv:Equivalence Rg:D -> DT':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y' : B, T x0 y' /\ g (x0, y') = (x0, y0)):A -> B -> Propx:Ay:BHy:T x yHeq1: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:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> PropHequiv:Equivalence Rg:D -> DT':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y'0 : B, T x0 y'0 /\ g (x0, y'0) = (x0, y0)):A -> B -> Propx:Ay:BHy:T x yx':Ay':BHeq1: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'0Heq2:snd (x, y) = snd (x', y')T x y /\ (x', y') = (x, snd (x', y'))subst; easy.ReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> PropHequiv:Equivalence Rg:D -> DT':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y'0 : B, T x0 y'0 /\ g (x0, y'0) = (x0, y0)):A -> B -> Propx:Ay:BHy:T x yx':Ay':BHeq1: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'0Heq2:y = y'T x y /\ (x', y') = (x, snd (x', y'))ReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> PropHequiv:Equivalence Rg:D -> DT':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y' : B, T x0 y' /\ g (x0, y') = (x0, y0)):A -> B -> Propx:Ay:BHy:T x yHeq1: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:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> PropHequiv:Equivalence Rg:D -> DT':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y'0 : B, T x0 y'0 /\ g (x0, y'0) = (x0, y0)):A -> B -> Propx:Ay:BHy:T x yHeq1: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':BHy':T x y'y'':BHy'':T x y''Heq:g (x, y'') = (x, y')snd (g (x, y)) = y'ReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> PropHequiv:Equivalence Rg:D -> DT':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y'0 : B, T x0 y'0 /\ g (x0, y'0) = (x0, y0)):A -> B -> Propx:Ay:BHy:T x yHeq1: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':BHy':T x y'y'':BHy'':T x y''Heq:g (x, y'') = (x, y')snd (x, y') = y'ReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> PropHequiv:Equivalence Rg:D -> DT':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y'0 : B, T x0 y'0 /\ g (x0, y'0) = (x0, y0)):A -> B -> Propx:Ay:BHy:T x yHeq1: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':BHy':T x y'y'':BHy'':T x y''Heq:g (x, y'') = (x, y')R (x, y) (x, y'')ReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> PropHequiv:Equivalence Rg:D -> DT':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y'0 : B, T x0 y'0 /\ g (x0, y'0) = (x0, y0)):A -> B -> Propx:Ay:BHy:T x yHeq1: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':BHy':T x y'y'':BHy'':T x y''Heq:g (x, y'') = (x, y')R (x, y) (x, y'')ReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> PropHequiv:Equivalence Rg:D -> DT':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y'0 : B, T x0 y'0 /\ g (x0, y'0) = (x0, y0)):A -> B -> Propx:Ay:BHy:T x yHeq1: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':BHy':T x y'y'':BHy'':T x y''Heq:g (x, y'') = (x, y')fst (x, y) = fst (x, y'')ReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> PropHequiv:Equivalence Rg:D -> DT':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y'0 : B, T x0 y'0 /\ g (x0, y'0) = (x0, y0)):A -> B -> Propx:Ay:BHy:T x yHeq1: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':BHy':T x y'y'':BHy'':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:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> PropHequiv:Equivalence Rg:D -> DT':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y'0 : B, T x0 y'0 /\ g (x0, y'0) = (x0, y0)):A -> B -> Propx:Ay:BHy:T x yHeq1: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':BHy':T x y'y'':BHy'':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:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> PropHequiv:Equivalence Rg:D -> DT':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y' : B, T x0 y' /\ g (x0, y') = (x0, y0)):A -> B -> Propx:Ay:BHy:T x yHeq1: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:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> PropHequiv:Equivalence Rg:D -> DT':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y' : B, T x0 y' /\ g (x0, y') = (x0, y0)):A -> B -> Propx:Ay:BHy:T x yHeq1: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:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> PropHequiv:Equivalence Rg:D -> DT':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y' : B, T x0 y' /\ g (x0, y') = (x0, y0)):A -> B -> Propx:Ay:BHy:T x yHeq1: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:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> PropHequiv:Equivalence Rg:D -> DT':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y' : B, T x0 y' /\ g (x0, y') = (x0, y0)):A -> B -> Propx:Ay:BHy:T x yHeq1: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:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> PropHequiv:Equivalence Rg:D -> DT':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y' : B, T x0 y' /\ g (x0, y') = (x0, y0)):A -> B -> Propx:Ay:BHy:T x yHeq1: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:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> PropHequiv:Equivalence Rg:D -> DT':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y' : B, T x0 y' /\ g (x0, y') = (x0, y0)):A -> B -> Propx:Ay:BHy:T x yHeq1: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'assumption.ReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> PropHequiv:Equivalence Rg:D -> DT':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y' : B, T x0 y' /\ g (x0, y') = (x0, y0)):A -> B -> Propx:Ay:BHy:T x yHeq1: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:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> PropHequiv:Equivalence Rg:D -> DT':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y' : B, T x0 y' /\ g (x0, y') = (x0, y0)):A -> B -> Propx:Ay:BHy:T x yHeq1: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:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> PropHequiv:Equivalence Rg:D -> DT':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y' : B, T x0 y' /\ g (x0, y') = (x0, y0)):A -> B -> Propx:Ay:BHy:T x yHeq1: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:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> PropHequiv:Equivalence Rg:D -> DT':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y'0 : B, T x0 y'0 /\ g (x0, y'0) = (x0, y0)):A -> B -> Propx:Ay:BHy:T x yx':Ay':BHeq1: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'0Hgy:T (fst (x, y)) (snd (x', y'))T x y /\ (x', y') = (x, snd (x', y'))subst x'; easy.ReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> PropHequiv:Equivalence Rg:D -> DT':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y'0 : B, T x0 y'0 /\ g (x0, y'0) = (x0, y0)):A -> B -> Propx:Ay:BHy:T x yx':Ay':BHeq1: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'0Hgy:T (fst (x, y)) (snd (x', y'))T x y /\ (x', y') = (x, snd (x', y'))ReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> PropHequiv:Equivalence Rg:D -> DT':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y' : B, T x0 y' /\ g (x0, y') = (x0, y0)):A -> B -> Propx:Ay:BHy:T x yHeq1: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:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> PropHequiv:Equivalence Rg:D -> DT':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y'0 : B, T x0 y'0 /\ g (x0, y'0) = (x0, y0)):A -> B -> Propx:Ay:BHy:T x yHeq1: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':BHy':T x y'y'':BHy'':T x y''Heq:g (x, y'') = (x, y')snd (g (x, y)) = y'ReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> PropHequiv:Equivalence Rg:D -> DT':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y'0 : B, T x0 y'0 /\ g (x0, y'0) = (x0, y0)):A -> B -> Propx:Ay:BHy:T x yHeq1: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':BHy':T x y'y'':BHy'':T x y''Heq:g (x, y'') = (x, y')snd (x, y') = y'ReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> PropHequiv:Equivalence Rg:D -> DT':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y'0 : B, T x0 y'0 /\ g (x0, y'0) = (x0, y0)):A -> B -> Propx:Ay:BHy:T x yHeq1: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':BHy':T x y'y'':BHy'':T x y''Heq:g (x, y'') = (x, y')R (x, y) (x, y'')ReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> PropHequiv:Equivalence Rg:D -> DT':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y'0 : B, T x0 y'0 /\ g (x0, y'0) = (x0, y0)):A -> B -> Propx:Ay:BHy:T x yHeq1: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':BHy':T x y'y'':BHy'':T x y''Heq:g (x, y'') = (x, y')R (x, y) (x, y'')ReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> PropHequiv:Equivalence Rg:D -> DT':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y'0 : B, T x0 y'0 /\ g (x0, y'0) = (x0, y0)):A -> B -> Propx:Ay:BHy:T x yHeq1: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':BHy':T x y'y'':BHy'':T x y''Heq:g (x, y'') = (x, y')fst (x, y) = fst (x, y'')ReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> PropHequiv:Equivalence Rg:D -> DT':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y'0 : B, T x0 y'0 /\ g (x0, y'0) = (x0, y0)):A -> B -> Propx:Ay:BHy:T x yHeq1: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':BHy':T x y'y'':BHy'':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. (**********************************************************************)ReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeT:A -> B -> PropHexists:forall x0 : A, exists y0 : B, T x0 y0D:=(A * B)%type:TypeR:=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 -> PropHequiv:Equivalence Rg:D -> DT':=fun (x0 : A) (y0 : B) => T x0 y0 /\ (exists y'0 : B, T x0 y'0 /\ g (x0, y'0) = (x0, y0)):A -> B -> Propx:Ay:BHy:T x yHeq1: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':BHy':T x y'y'':BHy'':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)))
forall A B : Type, GeneralizedSetoidFunctionalChoice_on A B -> SetoidFunctionalChoice_on A Bforall A B : Type, GeneralizedSetoidFunctionalChoice_on A B -> SetoidFunctionalChoice_on A BA, B:TypeGenSetoidFunChoice:GeneralizedSetoidFunctionalChoice_on A BR:A -> A -> PropT:A -> B -> PropHequiv:Equivalence RHcompat:forall (x x' : A) (y : B), R x x' -> T x y -> T x' yHex:forall x : A, exists y : B, T x yexists f : A -> B, forall x : A, T x (f x) /\ (forall x' : A, R x x' -> f x = f x')A, B:TypeGenSetoidFunChoice:GeneralizedSetoidFunctionalChoice_on A BR:A -> A -> PropT:A -> B -> PropHequiv:Equivalence RHcompat:forall (x x' : A) (y : B), R x x' -> T x y -> T x' yHex:forall x : A, exists y : B, T x yEquivalence eqA, B:TypeGenSetoidFunChoice:GeneralizedSetoidFunctionalChoice_on A BR:A -> A -> PropT:A -> B -> PropHequiv:Equivalence RHcompat:forall (x x' : A) (y : B), R x x' -> T x y -> T x' yHex:forall x : A, exists y : B, T x yforall (x x' : A) (y y' : B), R x x' -> y = y' -> T x y -> T x' y'A, B:TypeGenSetoidFunChoice:GeneralizedSetoidFunctionalChoice_on A BR:A -> A -> PropT:A -> B -> PropHequiv:Equivalence RHcompat:forall (x x' : A) (y : B), R x x' -> T x y -> T x' yHex:forall x : A, exists y : B, T x yforall (x x' : A) (y y' : B), R x x' -> y = y' -> T x y -> T x' y'firstorder. Qed.A, B:TypeGenSetoidFunChoice:GeneralizedSetoidFunctionalChoice_on A BR:A -> A -> PropT:A -> B -> PropHequiv:Equivalence RHcompat:forall (x0 x'0 : A) (y0 : B), R x0 x'0 -> T x0 y0 -> T x'0 y0Hex:forall x0 : A, exists y0 : B, T x0 y0x, x':Ay:BH:R x x'T x y -> T x' yforall A B : Type, SetoidFunctionalChoice_on A B -> GeneralizedSetoidFunctionalChoice_on A Bforall A B : Type, SetoidFunctionalChoice_on A B -> GeneralizedSetoidFunctionalChoice_on A BA, B:TypeSetoidFunChoice:SetoidFunctionalChoice_on A BR:A -> A -> PropS:B -> B -> PropT:A -> B -> PropHequivR:Equivalence RHequivS:Equivalence SHcompat: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 yexists f : A -> B, forall x : A, T x (f x) /\ (forall x' : A, R x x' -> S (f x) (f x'))A, B:TypeSetoidFunChoice:SetoidFunctionalChoice_on A BR:A -> A -> PropS:B -> B -> PropT:A -> B -> PropHequivR:Equivalence RHequivS:Equivalence SHcompat: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 yforall (x x' : A) (y : B), R x x' -> T x y -> T x' yA, B:TypeSetoidFunChoice:SetoidFunctionalChoice_on A BR:A -> A -> PropS:B -> B -> PropT:A -> B -> PropHequivR:Equivalence RHequivS:Equivalence SHcompat: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 yf:A -> BHf: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'))intros; apply (Hcompat x x' y y); try easy.A, B:TypeSetoidFunChoice:SetoidFunctionalChoice_on A BR:A -> A -> PropS:B -> B -> PropT:A -> B -> PropHequivR:Equivalence RHequivS:Equivalence SHcompat: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 yforall (x x' : A) (y : B), R x x' -> T x y -> T x' yA, B:TypeSetoidFunChoice:SetoidFunctionalChoice_on A BR:A -> A -> PropS:B -> B -> PropT:A -> B -> PropHequivR:Equivalence RHequivS:Equivalence SHcompat: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 yf:A -> BHf: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:TypeSetoidFunChoice:SetoidFunctionalChoice_on A BR:A -> A -> PropS:B -> B -> PropT:A -> B -> PropHequivR:Equivalence RHequivS:Equivalence SHcompat: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 yf:A -> BHf: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:TypeSetoidFunChoice:SetoidFunctionalChoice_on A BR:A -> A -> PropS:B -> B -> PropT:A -> B -> PropHequivR:Equivalence RHequivS:Equivalence SHcompat: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 yf:A -> Bx:AHf: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'))now erewrite Huniq. Qed.A, B:TypeSetoidFunChoice:SetoidFunctionalChoice_on A BR:A -> A -> PropS:B -> B -> PropT:A -> B -> PropHequivR:Equivalence RHequivS:Equivalence SHcompat: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 yf:A -> Bx:AHf:T x (f x)Huniq:forall x'0 : A, R x x'0 -> f x = f x'0x':AH:R x x'S (f x) (f x')forall A B : Type, SetoidFunctionalChoice_on A B <-> GeneralizedSetoidFunctionalChoice_on A Bsplit; 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 <-> GeneralizedSetoidFunctionalChoice_on A Bforall A B : Type, SetoidFunctionalChoice_on A B -> SimpleSetoidFunctionalChoice_on A Bforall A B : Type, SetoidFunctionalChoice_on A B -> SimpleSetoidFunctionalChoice_on A BA, B:TypeSetoidFunChoice:SetoidFunctionalChoice_on A BR:A -> A -> PropT:A -> B -> PropHequiv:Equivalence RHexists:forall x : A, exists y : B, forall x' : A, R x x' -> T x' yexists f : A -> B, forall x : A, T x (f x) /\ (forall x' : A, R x x' -> f x = f x')A, B:TypeSetoidFunChoice:SetoidFunctionalChoice_on A BR:A -> A -> PropT:A -> B -> PropHequiv:Equivalence RHexists:forall x : A, exists y : B, forall x' : A, R x x' -> T x' yT':=fun (x : A) (y : B) => forall x' : A, R x x' -> T x' y:A -> B -> Propexists f : A -> B, forall x : A, T x (f x) /\ (forall x' : A, R x x' -> f x = f x')A, B:TypeSetoidFunChoice:SetoidFunctionalChoice_on A BR:A -> A -> PropT:A -> B -> PropHequiv:Equivalence RHexists:forall x : A, exists y : B, forall x' : A, R x x' -> T x' yT':=fun (x : A) (y : B) => forall x' : A, R x x' -> T x' y:A -> B -> PropHcompat:forall (x x' : A) (y : B), R x x' -> T' x y -> T' x' yexists f : A -> B, forall x : A, T x (f x) /\ (forall x' : A, R x x' -> f x = f x')A, B:TypeSetoidFunChoice:SetoidFunctionalChoice_on A BR:A -> A -> PropT:A -> B -> PropHequiv:Equivalence RHexists:forall x : A, exists y : B, forall x' : A, R x x' -> T x' yT':=fun (x : A) (y : B) => forall x' : A, R x x' -> T x' y:A -> B -> PropHcompat:forall (x x' : A) (y : B), R x x' -> T' x y -> T' x' yf:A -> BHf: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')firstorder. Qed.A, B:TypeSetoidFunChoice:SetoidFunctionalChoice_on A BR:A -> A -> PropT:A -> B -> PropHequiv:Equivalence RHexists:forall x : A, exists y : B, forall x' : A, R x x' -> T x' yT':=fun (x : A) (y : B) => forall x' : A, R x x' -> T x' y:A -> B -> PropHcompat:forall (x x' : A) (y : B), R x x' -> T' x y -> T' x' yf:A -> BHf: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')forall A B : Type, SimpleSetoidFunctionalChoice_on A B -> SetoidFunctionalChoice_on A Bforall A B : Type, SimpleSetoidFunctionalChoice_on A B -> SetoidFunctionalChoice_on A Bdestruct (SimpleSetoidFunChoice R T Hequiv) as (f,Hf); firstorder. Qed.A, B:TypeSimpleSetoidFunChoice:SimpleSetoidFunctionalChoice_on A BR:A -> A -> PropT:A -> B -> PropHequiv:Equivalence RHcompat:forall (x x' : A) (y : B), R x x' -> T x y -> T x' yHexists:forall x : A, exists y : B, T x yexists f : A -> B, forall x : A, T x (f x) /\ (forall x' : A, R x x' -> f x = f x')forall A B : Type, SetoidFunctionalChoice_on A B <-> SimpleSetoidFunctionalChoice_on A Bsplit; auto using simple_setoid_fun_choice_imp_setoid_fun_choice, setoid_fun_choice_imp_simple_setoid_fun_choice. Qed. (**********************************************************************)forall A B : Type, SetoidFunctionalChoice_on A B <-> SimpleSetoidFunctionalChoice_on A B
forall A B : Type, SetoidFunctionalChoice_on A B -> FunctionalChoice_on A Bforall A B : Type, SetoidFunctionalChoice_on A B -> FunctionalChoice_on A BA, B:TypeSetoidFunChoice:SetoidFunctionalChoice_on A BT:A -> B -> PropHexists:forall x : A, exists y : B, T x yexists f : A -> B, forall x : A, T x (f x)A, B:TypeSetoidFunChoice:SetoidFunctionalChoice_on A BT:A -> B -> PropHexists:forall x : A, exists y : B, T x yEquivalence eqA, B:TypeSetoidFunChoice:SetoidFunctionalChoice_on A BT:A -> B -> PropHexists:forall x : A, exists y : B, T x yforall (x x' : A) (y : B), x = x' -> T x y -> T x' yA, B:TypeSetoidFunChoice:SetoidFunctionalChoice_on A BT:A -> B -> PropHexists:forall x : A, exists y : B, T x yforall x : A, exists y : B, T x yA, B:TypeSetoidFunChoice:SetoidFunctionalChoice_on A BT:A -> B -> PropHexists:forall x : A, exists y : B, T x yf:A -> BHf: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)apply eq_equivalence.A, B:TypeSetoidFunChoice:SetoidFunctionalChoice_on A BT:A -> B -> PropHexists:forall x : A, exists y : B, T x yEquivalence eqnow intros * ->.A, B:TypeSetoidFunChoice:SetoidFunctionalChoice_on A BT:A -> B -> PropHexists:forall x : A, exists y : B, T x yforall (x x' : A) (y : B), x = x' -> T x y -> T x' yassumption.A, B:TypeSetoidFunChoice:SetoidFunctionalChoice_on A BT:A -> B -> PropHexists:forall x : A, exists y : B, T x yforall x : A, exists y : B, T x yA, B:TypeSetoidFunChoice:SetoidFunctionalChoice_on A BT:A -> B -> PropHexists:forall x : A, exists y : B, T x yf:A -> BHf: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)firstorder. Qed.A, B:TypeSetoidFunChoice:SetoidFunctionalChoice_on A BT:A -> B -> PropHexists:forall x : A, exists y : B, T x yf:A -> BHf:forall x : A, T x (f x) /\ (forall x' : A, x = x' -> f x = f x')forall x : A, T x (f x)forall A B : Type, SetoidFunctionalChoice_on A B -> FunctionalRelReification_on A Bforall A B : Type, SetoidFunctionalChoice_on A B -> FunctionalRelReification_on A BA, B:TypeSetoidFunChoice:SetoidFunctionalChoice_on A BFunctionalRelReification_on A Bnow apply setoid_fun_choice_imp_fun_choice. Qed.A, B:TypeSetoidFunChoice:SetoidFunctionalChoice_on A BFunctionalChoice_on A BSetoidFunctionalChoice -> RepresentativeFunctionalChoiceSetoidFunctionalChoice -> RepresentativeFunctionalChoiceapply SetoidFunChoice; firstorder. Qed.SetoidFunChoice:SetoidFunctionalChoiceA:TypeR:A -> A -> PropHequiv:Equivalence Rexists f : A -> A, forall x : A, R x (f x) /\ (forall x' : A, R x x' -> f x = f x')FunctionalRelReification -> RepresentativeFunctionalChoice -> SetoidFunctionalChoiceFunctionalRelReification -> RepresentativeFunctionalChoice -> SetoidFunctionalChoiceFunRelReify:FunctionalRelReificationReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeR:A -> A -> PropT:A -> B -> PropHequiv:Equivalence RHcompat:forall (x x' : A) (y : B), R x x' -> T x y -> T x' yHexists:forall x : A, exists y : B, T x yexists f : A -> B, forall x : A, T x (f x) /\ (forall x' : A, R x x' -> f x = f x')FunRelReify:FunctionalRelReificationReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeR:A -> A -> PropT:A -> B -> PropHequiv:Equivalence RHcompat:forall (x x' : A) (y : B), R x x' -> T x y -> T x' yHexists:forall x : A, exists y : B, T x yFunctionalChoiceFunRelReify:FunctionalRelReificationReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeR:A -> A -> PropT:A -> B -> PropHequiv:Equivalence RHcompat:forall (x x' : A) (y : B), R x x' -> T x y -> T x' yHexists:forall x : A, exists y : B, T x yFunChoice:FunctionalChoiceexists f : A -> B, forall x : A, T x (f x) /\ (forall x' : A, R x x' -> f x = f x')FunRelReify:FunctionalRelReificationReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeR:A -> A -> PropT:A -> B -> PropHequiv:Equivalence RHcompat:forall (x x' : A) (y : B), R x x' -> T x y -> T x' yHexists:forall x : A, exists y : B, T x yFunctionalChoiceFunRelReify:FunctionalRelReificationReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeR:A -> A -> PropT:A -> B -> PropHequiv:Equivalence RHcompat:forall (x x' : A) (y : B), R x x' -> T x y -> T x' yHexists:forall x : A, exists y : B, T x yA', B':TypeFunctionalChoice_on A' B'FunRelReify:FunctionalRelReificationReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeR:A -> A -> PropT:A -> B -> PropHequiv:Equivalence RHcompat:forall (x x' : A) (y : B), R x x' -> T x y -> T x' yHexists:forall x : A, exists y : B, T x yA', B':TypeFunctionalRelReification_on A' B'FunRelReify:FunctionalRelReificationReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeR:A -> A -> PropT:A -> B -> PropHequiv:Equivalence RHcompat:forall (x x' : A) (y : B), R x x' -> T x y -> T x' yHexists:forall x : A, exists y : B, T x yA', B':TypeRelationalChoice_on A' B'apply FunRelReify.FunRelReify:FunctionalRelReificationReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeR:A -> A -> PropT:A -> B -> PropHequiv:Equivalence RHcompat:forall (x x' : A) (y : B), R x x' -> T x y -> T x' yHexists:forall x : A, exists y : B, T x yA', B':TypeFunctionalRelReification_on A' B'now apply repr_fun_choice_imp_relational_choice.FunRelReify:FunctionalRelReificationReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeR:A -> A -> PropT:A -> B -> PropHequiv:Equivalence RHcompat:forall (x x' : A) (y : B), R x x' -> T x y -> T x' yHexists:forall x : A, exists y : B, T x yA', B':TypeRelationalChoice_on A' B'FunRelReify:FunctionalRelReificationReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeR:A -> A -> PropT:A -> B -> PropHequiv:Equivalence RHcompat:forall (x x' : A) (y : B), R x x' -> T x y -> T x' yHexists:forall x : A, exists y : B, T x yFunChoice:FunctionalChoiceexists f : A -> B, forall x : A, T x (f x) /\ (forall x' : A, R x x' -> f x = f x')FunRelReify:FunctionalRelReificationReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeR:A -> A -> PropT:A -> B -> PropHequiv:Equivalence RHcompat:forall (x x' : A) (y : B), R x x' -> T x y -> T x' yHexists:forall x : A, exists y : B, T x yFunChoice:FunctionalChoicef:A -> BHf: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:FunctionalRelReificationReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeR:A -> A -> PropT:A -> B -> PropHequiv:Equivalence RHcompat:forall (x x' : A) (y : B), R x x' -> T x y -> T x' yHexists:forall x : A, exists y : B, T x yFunChoice:FunctionalChoicef:A -> BHf:forall x : A, T x (f x)g:A -> AHg: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:FunctionalRelReificationReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeR:A -> A -> PropT:A -> B -> PropHequiv:Equivalence RHcompat:forall (x x' : A) (y : B), R x x' -> T x y -> T x' yHexists:forall x : A, exists y : B, T x yFunChoice:FunctionalChoicef:A -> BHf:forall x : A, T x (f x)g:A -> AHg: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:FunctionalRelReificationReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeR:A -> A -> PropT:A -> B -> PropHequiv:Equivalence RHcompat:forall (x0 x' : A) (y : B), R x0 x' -> T x0 y -> T x' yHexists:forall x0 : A, exists y : B, T x0 yFunChoice:FunctionalChoicef:A -> BHf:forall x0 : A, T x0 (f x0)g:A -> AHg:forall x0 : A, R x0 (g x0) /\ (forall x' : A, R x0 x' -> g x0 = g x')x:AT x (f (g x)) /\ (forall x' : A, R x x' -> f (g x) = f (g x'))FunRelReify:FunctionalRelReificationReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeR:A -> A -> PropT:A -> B -> PropHequiv:Equivalence RHcompat:forall (x0 x' : A) (y : B), R x0 x' -> T x0 y -> T x' yHexists:forall x0 : A, exists y : B, T x0 yFunChoice:FunctionalChoicef:A -> BHf:forall x0 : A, T x0 (f x0)g:A -> AHg:forall x0 : A, R x0 (g x0) /\ (forall x' : A, R x0 x' -> g x0 = g x')x:AHgx: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:FunctionalRelReificationReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeR:A -> A -> PropT:A -> B -> PropHequiv:Equivalence RHcompat:forall (x0 x' : A) (y : B), R x0 x' -> T x0 y -> T x' yHexists:forall x0 : A, exists y : B, T x0 yFunChoice:FunctionalChoicef:A -> BHf:forall x0 : A, T x0 (f x0)g:A -> AHg:forall x0 : A, R x0 (g x0) /\ (forall x' : A, R x0 x' -> g x0 = g x')x:AHgx:R x (g x)HRuniq:forall x' : A, R x x' -> g x = g x'T x (f (g x))FunRelReify:FunctionalRelReificationReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeR:A -> A -> PropT:A -> B -> PropHequiv:Equivalence RHcompat:forall (x0 x' : A) (y : B), R x0 x' -> T x0 y -> T x' yHexists:forall x0 : A, exists y : B, T x0 yFunChoice:FunctionalChoicef:A -> BHf:forall x0 : A, T x0 (f x0)g:A -> AHg:forall x0 : A, R x0 (g x0) /\ (forall x' : A, R x0 x' -> g x0 = g x')x:AHgx: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:FunctionalRelReificationReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeR:A -> A -> PropT:A -> B -> PropHequiv:Equivalence RHcompat:forall (x0 x' : A) (y : B), R x0 x' -> T x0 y -> T x' yHexists:forall x0 : A, exists y : B, T x0 yFunChoice:FunctionalChoicef:A -> BHf:forall x0 : A, T x0 (f x0)g:A -> AHg:forall x0 : A, R x0 (g x0) /\ (forall x' : A, R x0 x' -> g x0 = g x')x:AHgx:R x (g x)HRuniq:forall x' : A, R x x' -> g x = g x'T x (f (g x))FunRelReify:FunctionalRelReificationReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeR:A -> A -> PropT:A -> B -> PropHequiv:Equivalence RHcompat:forall (x0 x' : A) (y : B), R x0 x' -> T x0 y -> T x' yHexists:forall x0 : A, exists y : B, T x0 yFunChoice:FunctionalChoicef:A -> BHf:forall x0 : A, T x0 (f x0)g:A -> AHg:forall x0 : A, R x0 (g x0) /\ (forall x' : A, R x0 x' -> g x0 = g x')x:AHgx:R x (g x)HRuniq:forall x' : A, R x x' -> g x = g x'R ?x xFunRelReify:FunctionalRelReificationReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeR:A -> A -> PropT:A -> B -> PropHequiv:Equivalence RHcompat:forall (x0 x' : A) (y : B), R x0 x' -> T x0 y -> T x' yHexists:forall x0 : A, exists y : B, T x0 yFunChoice:FunctionalChoicef:A -> BHf:forall x0 : A, T x0 (f x0)g:A -> AHg:forall x0 : A, R x0 (g x0) /\ (forall x' : A, R x0 x' -> g x0 = g x')x:AHgx:R x (g x)HRuniq:forall x' : A, R x x' -> g x = g x'T ?x (f (g x))FunRelReify:FunctionalRelReificationReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeR:A -> A -> PropT:A -> B -> PropHequiv:Equivalence RHcompat:forall (x0 x' : A) (y : B), R x0 x' -> T x0 y -> T x' yHexists:forall x0 : A, exists y : B, T x0 yFunChoice:FunctionalChoicef:A -> BHf:forall x0 : A, T x0 (f x0)g:A -> AHg:forall x0 : A, R x0 (g x0) /\ (forall x' : A, R x0 x' -> g x0 = g x')x:AHgx:R x (g x)HRuniq:forall x' : A, R x x' -> g x = g x'R x ?xFunRelReify:FunctionalRelReificationReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeR:A -> A -> PropT:A -> B -> PropHequiv:Equivalence RHcompat:forall (x0 x' : A) (y : B), R x0 x' -> T x0 y -> T x' yHexists:forall x0 : A, exists y : B, T x0 yFunChoice:FunctionalChoicef:A -> BHf:forall x0 : A, T x0 (f x0)g:A -> AHg:forall x0 : A, R x0 (g x0) /\ (forall x' : A, R x0 x' -> g x0 = g x')x:AHgx:R x (g x)HRuniq:forall x' : A, R x x' -> g x = g x'T ?x (f (g x))apply Hf.FunRelReify:FunctionalRelReificationReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeR:A -> A -> PropT:A -> B -> PropHequiv:Equivalence RHcompat:forall (x0 x' : A) (y : B), R x0 x' -> T x0 y -> T x' yHexists:forall x0 : A, exists y : B, T x0 yFunChoice:FunctionalChoicef:A -> BHf:forall x0 : A, T x0 (f x0)g:A -> AHg:forall x0 : A, R x0 (g x0) /\ (forall x' : A, R x0 x' -> g x0 = g x')x:AHgx:R x (g x)HRuniq:forall x' : A, R x x' -> g x = g x'T (g x) (f (g x))FunRelReify:FunctionalRelReificationReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeR:A -> A -> PropT:A -> B -> PropHequiv:Equivalence RHcompat:forall (x0 x' : A) (y : B), R x0 x' -> T x0 y -> T x' yHexists:forall x0 : A, exists y : B, T x0 yFunChoice:FunctionalChoicef:A -> BHf:forall x0 : A, T x0 (f x0)g:A -> AHg:forall x0 : A, R x0 (g x0) /\ (forall x' : A, R x0 x' -> g x0 = g x')x:AHgx: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:FunctionalRelReificationReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeR:A -> A -> PropT:A -> B -> PropHequiv:Equivalence RHcompat:forall (x0 x' : A) (y0 : B), R x0 x' -> T x0 y0 -> T x' y0Hexists:forall x0 : A, exists y0 : B, T x0 y0FunChoice:FunctionalChoicef:A -> BHf:forall x0 : A, T x0 (f x0)g:A -> AHg:forall x0 : A, R x0 (g x0) /\ (forall x' : A, R x0 x' -> g x0 = g x')x:AHgx:R x (g x)HRuniq:forall x' : A, R x x' -> g x = g x'y:AHxy:R x yf (g x) = f (g y)auto. Qed.FunRelReify:FunctionalRelReificationReprFunChoice:RepresentativeFunctionalChoiceA, B:TypeR:A -> A -> PropT:A -> B -> PropHequiv:Equivalence RHcompat:forall (x0 x' : A) (y0 : B), R x0 x' -> T x0 y0 -> T x' y0Hexists:forall x0 : A, exists y0 : B, T x0 y0FunChoice:FunctionalChoicef:A -> BHf:forall x0 : A, T x0 (f x0)g:A -> AHg:forall x0 : A, R x0 (g x0) /\ (forall x' : A, R x0 x' -> g x0 = g x')x:AHgx:R x (g x)HRuniq:forall x' : A, R x x' -> g x = g x'y:AHxy:R x yg x = g yFunctionalRelReification /\ RepresentativeFunctionalChoice <-> SetoidFunctionalChoiceFunctionalRelReification /\ RepresentativeFunctionalChoice <-> SetoidFunctionalChoiceH:FunctionalRelReification /\ RepresentativeFunctionalChoiceA, B:TypeSetoidFunctionalChoice_on A BH:SetoidFunctionalChoiceFunctionalRelReification /\ RepresentativeFunctionalChoicenow apply functional_rel_reification_and_repr_fun_choice_imp_setoid_fun_choice.H:FunctionalRelReification /\ RepresentativeFunctionalChoiceA, B:TypeSetoidFunctionalChoice_on A BH:SetoidFunctionalChoiceFunctionalRelReification /\ RepresentativeFunctionalChoiceH:SetoidFunctionalChoiceFunctionalRelReificationH:SetoidFunctionalChoiceRepresentativeFunctionalChoicenow intros A B; apply setoid_fun_choice_imp_functional_rel_reification.H:SetoidFunctionalChoiceFunctionalRelReificationnow apply setoid_fun_choice_imp_repr_fun_choice. Qed.H:SetoidFunctionalChoiceRepresentativeFunctionalChoice
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
(**********************************************************************)
Import EqNotations.
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 -> RepresentativeFunctionalChoiceFunctionalChoice -> ExtensionalFunctionRepresentative -> ExcludedMiddle -> RepresentativeFunctionalChoiceFunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive Rexists f : A -> A, forall x : A, R x (f x) /\ (forall x' : A, R x x' -> f x = f x')FunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive Rforall P : Prop, exists b : bool, b = true <-> PFunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pexists f : A -> A, forall x : A, R x (f x) /\ (forall x' : A, R x x' -> f x = f x')FunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive Rforall P : Prop, exists b : bool, b = true <-> PFunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RP:Propexists b : bool, b = true <-> PFunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RP:PropH:Pexists b : bool, b = true <-> PFunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RP:PropH:~ Pexists b : bool, b = true <-> Pexists true; firstorder.FunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RP:PropH:Pexists b : bool, b = true <-> Pexists false; easy.FunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RP:PropH:~ Pexists b : bool, b = true <-> PFunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pexists f : A -> A, forall x : A, R x (f x) /\ (forall x' : A, R x x' -> f x = f x')FunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pc:Prop -> boolHc:forall x : Prop, c x = true <-> xexists f : A -> A, forall x : A, R x (f x) /\ (forall x' : A, R x x' -> f x = f x')FunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pc:Prop -> boolHc:forall x : Prop, c x = true <-> xclass_of:=fun a y : A => c (R a y):A -> A -> boolexists f : A -> A, forall x : A, R x (f x) /\ (forall x' : A, R x x' -> f x = f x')FunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pc:Prop -> boolHc:forall x : Prop, c x = true <-> xclass_of:=fun a y : A => c (R a y):A -> A -> boolisclass:=fun f : A -> bool => exists x : A, f x = true:(A -> bool) -> Propexists f : A -> A, forall x : A, R x (f x) /\ (forall x' : A, R x x' -> f x = f x')FunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pc:Prop -> boolHc:forall x : Prop, c x = true <-> xclass_of:=fun a y : A => c (R a y):A -> A -> boolisclass:=fun f : A -> bool => exists x : A, f x = true:(A -> bool) -> Propclass:={f : A -> bool | isclass f}:Typeexists f : A -> A, forall x : A, R x (f x) /\ (forall x' : A, R x x' -> f x = f x')FunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pc:Prop -> boolHc:forall x : Prop, c x = true <-> xclass_of:=fun a y : A => c (R a y):A -> A -> boolisclass:=fun f : A -> bool => exists x : A, f x = true:(A -> bool) -> Propclass:={f : A -> bool | isclass f}:Typecontains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Propexists f : A -> A, forall x : A, R x (f x) /\ (forall x' : A, R x x' -> f x = f x')FunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pc:Prop -> boolHc:forall x : Prop, c x = true <-> xclass_of:=fun a y : A => c (R a y):A -> A -> boolisclass:=fun f : A -> bool => exists x : A, f x = true:(A -> bool) -> Propclass:={f : A -> bool | isclass f}:Typecontains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Propforall x : class, exists y : A, contains x yFunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pc:Prop -> boolHc:forall x : Prop, c x = true <-> xclass_of:=fun a y : A => c (R a y):A -> A -> boolisclass:=fun f0 : A -> bool => exists x : A, f0 x = true:(A -> bool) -> Propclass:={f0 : A -> bool | isclass f0}:Typecontains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Propf:class -> AHf: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:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pc:Prop -> boolHc:forall x : Prop, c x = true <-> xclass_of:=fun a y : A => c (R a y):A -> A -> boolisclass:=fun f : A -> bool => exists x : A, f x = true:(A -> bool) -> Propclass:={f : A -> bool | isclass f}:Typecontains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Propforall x : class, exists y : A, contains x yFunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pc:Prop -> boolHc:forall x : Prop, c x = true <-> xclass_of:=fun a y : A => c (R a y):A -> A -> boolisclass:=fun f0 : A -> bool => exists x : A, f0 x = true:(A -> bool) -> Propclass:={f0 : A -> bool | isclass f0}:Typecontains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Propf:classexists y : A, contains f yFunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pc:Prop -> boolHc:forall x0 : Prop, c x0 = true <-> x0class_of:=fun a y : A => c (R a y):A -> A -> boolisclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Propclass:={f0 : A -> bool | isclass f0}:Typecontains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Propf:classx:AHx:proj1_sig f x = trueexists y : A, contains f yeasy.FunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pc:Prop -> boolHc:forall x0 : Prop, c x0 = true <-> x0class_of:=fun a y : A => c (R a y):A -> A -> boolisclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Propclass:={f0 : A -> bool | isclass f0}:Typecontains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Propf:classx:AHx:proj1_sig f x = truecontains f xFunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pc:Prop -> boolHc:forall x : Prop, c x = true <-> xclass_of:=fun a y : A => c (R a y):A -> A -> boolisclass:=fun f0 : A -> bool => exists x : A, f0 x = true:(A -> bool) -> Propclass:={f0 : A -> bool | isclass f0}:Typecontains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Propf:class -> AHf: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:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pc:Prop -> boolHc:forall x : Prop, c x = true <-> xclass_of:=fun a y : A => c (R a y):A -> A -> boolisclass:=fun f0 : A -> bool => exists x : A, f0 x = true:(A -> bool) -> Propclass:={f0 : A -> bool | isclass f0}:Typecontains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Propf:class -> AHf:forall x : class, contains x (f x)h:(A -> bool) -> A -> boolHh: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:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pc:Prop -> boolHc:forall x : Prop, c x = true <-> xclass_of:=fun a y : A => c (R a y):A -> A -> boolisclass:=fun f0 : A -> bool => exists x : A, f0 x = true:(A -> bool) -> Propclass:={f0 : A -> bool | isclass f0}:Typecontains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Propf:class -> AHf:forall x : class, contains x (f x)h:(A -> bool) -> A -> boolHh: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:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pc:Prop -> boolHc:forall x : Prop, c x = true <-> xclass_of:=fun a y : A => c (R a y):A -> A -> boolisclass:=fun f0 : A -> bool => exists x : A, f0 x = true:(A -> bool) -> Propclass:={f0 : A -> bool | isclass f0}:Typecontains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Propf:class -> AHf:forall x : class, contains x (f x)h:(A -> bool) -> A -> boolHh: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:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pc:Prop -> boolHc:forall x : Prop, c x = true <-> xclass_of:=fun a y : A => c (R a y):A -> A -> boolisclass:=fun f0 : A -> bool => exists x : A, f0 x = true:(A -> bool) -> Propclass:={f0 : A -> bool | isclass f0}:Typecontains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Propf:class -> AHf:forall x : class, contains x (f x)h:(A -> bool) -> A -> boolHh: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:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pc:Prop -> boolHc:forall x : Prop, c x = true <-> xclass_of:=fun a0 y : A => c (R a0 y):A -> A -> boolisclass:=fun f0 : A -> bool => exists x : A, f0 x = true:(A -> bool) -> Propclass:={f0 : A -> bool | isclass f0}:Typecontains:=fun (c0 : class) (a0 : A) => proj1_sig c0 a0 = true:class -> A -> Propf:class -> AHf:forall x : class, contains x (f x)h:(A -> bool) -> A -> boolHh: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:Aisclass (h (class_of a))FunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pc:Prop -> boolHc:forall x : Prop, c x = true <-> xclass_of:=fun a0 y : A => c (R a0 y):A -> A -> boolisclass:=fun f0 : A -> bool => exists x : A, f0 x = true:(A -> bool) -> Propclass:={f0 : A -> bool | isclass f0}:Typecontains:=fun (c0 : class) (a0 : A) => proj1_sig c0 a0 = true:class -> A -> Propf:class -> AHf:forall x : class, contains x (f x)h:(A -> bool) -> A -> boolHh: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:Ah (class_of a) a = trueFunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pc:Prop -> boolHc:forall x : Prop, c x = true <-> xclass_of:=fun a0 y : A => c (R a0 y):A -> A -> boolisclass:=fun f0 : A -> bool => exists x : A, f0 x = true:(A -> bool) -> Propclass:={f0 : A -> bool | isclass f0}:Typecontains:=fun (c0 : class) (a0 : A) => proj1_sig c0 a0 = true:class -> A -> Propf:class -> AHf:forall x : class, contains x (f x)h:(A -> bool) -> A -> boolHh: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:AHa:forall x : A, class_of a x = h (class_of a) xHuniqa:forall g : A -> bool, (forall x : A, class_of a x = g x) -> h (class_of a) = h gh (class_of a) a = trueFunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pc:Prop -> boolHc:forall x : Prop, c x = true <-> xclass_of:=fun a0 y : A => c (R a0 y):A -> A -> boolisclass:=fun f0 : A -> bool => exists x : A, f0 x = true:(A -> bool) -> Propclass:={f0 : A -> bool | isclass f0}:Typecontains:=fun (c0 : class) (a0 : A) => proj1_sig c0 a0 = true:class -> A -> Propf:class -> AHf:forall x : class, contains x (f x)h:(A -> bool) -> A -> boolHh: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:AHa:forall x : A, class_of a x = h (class_of a) xHuniqa:forall g : A -> bool, (forall x : A, class_of a x = g x) -> h (class_of a) = h gclass_of a a = trueapply Hrefl.FunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pc:Prop -> boolHc:forall x : Prop, c x = true <-> xclass_of:=fun a0 y : A => c (R a0 y):A -> A -> boolisclass:=fun f0 : A -> bool => exists x : A, f0 x = true:(A -> bool) -> Propclass:={f0 : A -> bool | isclass f0}:Typecontains:=fun (c0 : class) (a0 : A) => proj1_sig c0 a0 = true:class -> A -> Propf:class -> AHf:forall x : class, contains x (f x)h:(A -> bool) -> A -> boolHh: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:AHa:forall x : A, class_of a x = h (class_of a) xHuniqa:forall g : A -> bool, (forall x : A, class_of a x = g x) -> h (class_of a) = h gR a aFunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pc:Prop -> boolHc:forall x : Prop, c x = true <-> xclass_of:=fun a y : A => c (R a y):A -> A -> boolisclass:=fun f0 : A -> bool => exists x : A, f0 x = true:(A -> bool) -> Propclass:={f0 : A -> bool | isclass f0}:Typecontains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Propf:class -> AHf:forall x : class, contains x (f x)h:(A -> bool) -> A -> boolHh: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:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pc:Prop -> boolHc:forall x : Prop, c x = true <-> xclass_of:=fun a y : A => c (R a y):A -> A -> boolisclass:=fun f0 : A -> bool => exists x : A, f0 x = true:(A -> bool) -> Propclass:={f0 : A -> bool | isclass f0}:Typecontains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Propf:class -> AHf:forall x : class, contains x (f x)h:(A -> bool) -> A -> boolHh: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 -> classexists f0 : A -> A, forall x : A, R x (f0 x) /\ (forall x' : A, R x x' -> f0 x = f0 x')FunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pc:Prop -> boolHc:forall x : Prop, c x = true <-> xclass_of:=fun a y : A => c (R a y):A -> A -> boolisclass:=fun f0 : A -> bool => exists x : A, f0 x = true:(A -> bool) -> Propclass:={f0 : A -> bool | isclass f0}:Typecontains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Propf:class -> AHf:forall x : class, contains x (f x)h:(A -> bool) -> A -> boolHh: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 -> classforall x : A, R x (f (f' x)) /\ (forall x' : A, R x x' -> f (f' x) = f (f' x'))FunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pc:Prop -> boolHc:forall x0 : Prop, c x0 = true <-> x0class_of:=fun a y : A => c (R a y):A -> A -> boolisclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Propclass:={f0 : A -> bool | isclass f0}:Typecontains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Propf:class -> AHf:forall x0 : class, contains x0 (f x0)h:(A -> bool) -> A -> boolHh: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 -> classx:AR x (f (f' x)) /\ (forall x' : A, R x x' -> f (f' x) = f (f' x'))FunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pc:Prop -> boolHc:forall x0 : Prop, c x0 = true <-> x0class_of:=fun a y : A => c (R a y):A -> A -> boolisclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Propclass:={f0 : A -> bool | isclass f0}:Typecontains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Propf:class -> AHf:forall x0 : class, contains x0 (f x0)h:(A -> bool) -> A -> boolHh: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 -> classx:AHx:forall x0 : A, class_of x x0 = h (class_of x) x0Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h gR x (f (f' x)) /\ (forall x' : A, R x x' -> f (f' x) = f (f' x'))FunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pc:Prop -> boolHc:forall x0 : Prop, c x0 = true <-> x0class_of:=fun a y : A => c (R a y):A -> A -> boolisclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Propclass:={f0 : A -> bool | isclass f0}:Typecontains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Propf:class -> AHf:forall x0 : class, contains x0 (f x0)h:(A -> bool) -> A -> boolHh: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 -> classx:AHx:forall x0 : A, class_of x x0 = h (class_of x) x0Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h gR x (f (f' x))FunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pc:Prop -> boolHc:forall x0 : Prop, c x0 = true <-> x0class_of:=fun a y : A => c (R a y):A -> A -> boolisclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Propclass:={f0 : A -> bool | isclass f0}:Typecontains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Propf:class -> AHf:forall x0 : class, contains x0 (f x0)h:(A -> bool) -> A -> boolHh: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 -> classx:AHx:forall x0 : A, class_of x x0 = h (class_of x) x0Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h gforall x' : A, R x x' -> f (f' x) = f (f' x')FunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pc:Prop -> boolHc:forall x0 : Prop, c x0 = true <-> x0class_of:=fun a y : A => c (R a y):A -> A -> boolisclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Propclass:={f0 : A -> bool | isclass f0}:Typecontains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Propf:class -> AHf:forall x0 : class, contains x0 (f x0)h:(A -> bool) -> A -> boolHh: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 -> classx:AHx:forall x0 : A, class_of x x0 = h (class_of x) x0Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h gR x (f (f' x))FunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pc:Prop -> boolHc:forall x0 : Prop, c x0 = true <-> x0class_of:=fun a y : A => c (R a y):A -> A -> boolisclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Propclass:={f0 : A -> bool | isclass f0}:Typecontains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Propf:class -> Ah:(A -> bool) -> A -> boolHisclass:forall a : A, isclass (h (class_of a))f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> classx:AHf: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) x0Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h gR x (f (f' x))FunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pc:Prop -> boolHc:forall x0 : Prop, c x0 = true <-> x0class_of:=fun a y : A => c (R a y):A -> A -> boolisclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Propclass:={f0 : A -> bool | isclass f0}:Typecontains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Propf:class -> Ah:(A -> bool) -> A -> boolHisclass:forall a : A, isclass (h (class_of a))f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> classx:AHf:proj1_sig (f' x) (f (f' x)) = trueHh: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) x0Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h gR x (f (f' x))FunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pc:Prop -> boolHc:forall x0 : Prop, c x0 = true <-> x0class_of:=fun a y : A => c (R a y):A -> A -> boolisclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Propclass:={f0 : A -> bool | isclass f0}:Typecontains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Propf:class -> Ah:(A -> bool) -> A -> boolHisclass:forall a : A, isclass (h (class_of a))f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> classx:AHf:h (class_of x) (f (f' x)) = trueHh: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) x0Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h gR x (f (f' x))FunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pc:Prop -> boolHc:forall x0 : Prop, c x0 = true <-> x0class_of:=fun a y : A => c (R a y):A -> A -> boolisclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Propclass:={f0 : A -> bool | isclass f0}:Typecontains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Propf:class -> Ah:(A -> bool) -> A -> boolHisclass:forall a : A, isclass (h (class_of a))f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> classx:AHf:class_of x (f (f' x)) = trueHh: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) x0Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h gR x (f (f' x))assumption.FunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pc:Prop -> boolHc:forall x0 : Prop, c x0 = true <-> x0class_of:=fun a y : A => c (R a y):A -> A -> boolisclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Propclass:={f0 : A -> bool | isclass f0}:Typecontains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Propf:class -> Ah:(A -> bool) -> A -> boolHisclass:forall a : A, isclass (h (class_of a))f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> classx:AHf:class_of x (f (f' x)) = trueHh: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) x0Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h gc (R x (f (f' x))) = trueFunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pc:Prop -> boolHc:forall x0 : Prop, c x0 = true <-> x0class_of:=fun a y : A => c (R a y):A -> A -> boolisclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Propclass:={f0 : A -> bool | isclass f0}:Typecontains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Propf:class -> AHf:forall x0 : class, contains x0 (f x0)h:(A -> bool) -> A -> boolHh: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 -> classx:AHx:forall x0 : A, class_of x x0 = h (class_of x) x0Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h gforall x' : A, R x x' -> f (f' x) = f (f' x')FunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pc:Prop -> boolHc:forall x0 : Prop, c x0 = true <-> x0class_of:=fun a y0 : A => c (R a y0):A -> A -> boolisclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Propclass:={f0 : A -> bool | isclass f0}:Typecontains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Propf:class -> AHf:forall x0 : class, contains x0 (f x0)h:(A -> bool) -> A -> boolHh: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 -> classx:AHx:forall x0 : A, class_of x x0 = h (class_of x) x0Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h gy:AHxy:R x yf (f' x) = f (f' y)FunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pc:Prop -> boolHc:forall x0 : Prop, c x0 = true <-> x0class_of:=fun a y0 : A => c (R a y0):A -> A -> boolisclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Propclass:={f0 : A -> bool | isclass f0}:Typecontains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Propf:class -> AHf:forall x0 : class, contains x0 (f x0)h:(A -> bool) -> A -> boolHh: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 -> classx:AHx:forall x0 : A, class_of x x0 = h (class_of x) x0Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h gy:AHxy:R x yf' x = f' yFunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pc:Prop -> boolHc:forall x0 : Prop, c x0 = true <-> x0class_of:=fun a y0 : A => c (R a y0):A -> A -> boolisclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Propclass:={f0 : A -> bool | isclass f0}:Typecontains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Propf:class -> AHf:forall x0 : class, contains x0 (f x0)h:(A -> bool) -> A -> boolHh: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 -> classx:AHx:forall x0 : A, class_of x x0 = h (class_of x) x0Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h gy:AHxy:R x yh (class_of x) = h (class_of y)FunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pc:Prop -> boolHc:forall x0 : Prop, c x0 = true <-> x0class_of:=fun a y0 : A => c (R a y0):A -> A -> boolisclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Propclass:={f0 : A -> bool | isclass f0}:Typecontains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Propf:class -> AHf:forall x0 : class, contains x0 (f x0)h:(A -> bool) -> A -> boolHh: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 -> classx:AHx:forall x0 : A, class_of x x0 = h (class_of x) x0Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h gy:AHxy:R x yHeq1:h (class_of x) = h (class_of y)f' x = f' yFunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pc:Prop -> boolHc:forall x0 : Prop, c x0 = true <-> x0class_of:=fun a y0 : A => c (R a y0):A -> A -> boolisclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Propclass:={f0 : A -> bool | isclass f0}:Typecontains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Propf:class -> AHf:forall x0 : class, contains x0 (f x0)h:(A -> bool) -> A -> boolHh: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 -> classx:AHx:forall x0 : A, class_of x x0 = h (class_of x) x0Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h gy:AHxy:R x yh (class_of x) = h (class_of y)FunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pc:Prop -> boolHc:forall x0 : Prop, c x0 = true <-> x0class_of:=fun a y0 : A => c (R a y0):A -> A -> boolisclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Propclass:={f0 : A -> bool | isclass f0}:Typecontains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Propf:class -> AHf:forall x0 : class, contains x0 (f x0)h:(A -> bool) -> A -> boolHh: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 -> classx:AHx:forall x0 : A, class_of x x0 = h (class_of x) x0Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h gy:AHxy:R x yforall x0 : A, class_of x x0 = class_of y x0FunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pc:Prop -> boolHc:forall x0 : Prop, c x0 = true <-> x0class_of:=fun a y0 : A => c (R a y0):A -> A -> boolisclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Propclass:={f0 : A -> bool | isclass f0}:Typecontains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Propf:class -> AHf:forall x0 : class, contains x0 (f x0)h:(A -> bool) -> A -> boolHh: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 -> classx:AHx:forall x0 : A, class_of x x0 = h (class_of x) x0Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h gy:AHxy:R x yz:Aclass_of x z = class_of y zFunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pc:Prop -> boolHc:forall x0 : Prop, c x0 = true <-> x0class_of:=fun a y0 : A => c (R a y0):A -> A -> boolisclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Propclass:={f0 : A -> bool | isclass f0}:Typecontains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Propf:class -> AHf:forall x0 : class, contains x0 (f x0)h:(A -> bool) -> A -> boolHh: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 -> classx:AHx:forall x0 : A, class_of x x0 = h (class_of x) x0Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h gy:AHxy:R x yz:Ac (R x z) = c (R y z)FunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pc:Prop -> boolHc:forall x0 : Prop, c x0 = true <-> x0class_of:=fun a y0 : A => c (R a y0):A -> A -> boolisclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Propclass:={f0 : A -> bool | isclass f0}:Typecontains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Propf:class -> AHf:forall x0 : class, contains x0 (f x0)h:(A -> bool) -> A -> boolHh: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 -> classx:AHx:forall x0 : A, class_of x x0 = h (class_of x) x0Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h gy:AHxy:R x yz:AHxz:c (R x z) = truetrue = c (R y z)FunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pc:Prop -> boolHc:forall x0 : Prop, c x0 = true <-> x0class_of:=fun a y0 : A => c (R a y0):A -> A -> boolisclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Propclass:={f0 : A -> bool | isclass f0}:Typecontains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Propf:class -> AHf:forall x0 : class, contains x0 (f x0)h:(A -> bool) -> A -> boolHh: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 -> classx:AHx:forall x0 : A, class_of x x0 = h (class_of x) x0Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h gy:AHxy:R x yz:AHxz:c (R x z) = falsefalse = c (R y z)FunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pc:Prop -> boolHc:forall x0 : Prop, c x0 = true <-> x0class_of:=fun a y0 : A => c (R a y0):A -> A -> boolisclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Propclass:={f0 : A -> bool | isclass f0}:Typecontains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Propf:class -> AHf:forall x0 : class, contains x0 (f x0)h:(A -> bool) -> A -> boolHh: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 -> classx:AHx:forall x0 : A, class_of x x0 = h (class_of x) x0Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h gy:AHxy:R x yz:AHxz:c (R x z) = truetrue = c (R y z)FunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pc:Prop -> boolHc:forall x0 : Prop, c x0 = true <-> x0class_of:=fun a y0 : A => c (R a y0):A -> A -> boolisclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Propclass:={f0 : A -> bool | isclass f0}:Typecontains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Propf:class -> AHf:forall x0 : class, contains x0 (f x0)h:(A -> bool) -> A -> boolHh: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 -> classx:AHx:forall x0 : A, class_of x x0 = h (class_of x) x0Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h gy:AHxy:R x yz:AHxz:c (R x z) = truec (R y z) = trueFunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pc:Prop -> boolHc:forall x0 : Prop, c x0 = true <-> x0class_of:=fun a y0 : A => c (R a y0):A -> A -> boolisclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Propclass:={f0 : A -> bool | isclass f0}:Typecontains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Propf:class -> AHf:forall x0 : class, contains x0 (f x0)h:(A -> bool) -> A -> boolHh: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 -> classx:AHx:forall x0 : A, class_of x x0 = h (class_of x) x0Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h gy:AHxy:R x yz:AHxz:c (R x z) = trueR y zfirstorder.FunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pc:Prop -> boolHc:forall x0 : Prop, c x0 = true <-> x0class_of:=fun a y0 : A => c (R a y0):A -> A -> boolisclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Propclass:={f0 : A -> bool | isclass f0}:Typecontains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Propf:class -> AHf:forall x0 : class, contains x0 (f x0)h:(A -> bool) -> A -> boolHh: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 -> classx:AHx:forall x0 : A, class_of x x0 = h (class_of x) x0Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h gy:AHxy:R x yz:AHxz:R x zR y zFunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pc:Prop -> boolHc:forall x0 : Prop, c x0 = true <-> x0class_of:=fun a y0 : A => c (R a y0):A -> A -> boolisclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Propclass:={f0 : A -> bool | isclass f0}:Typecontains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Propf:class -> AHf:forall x0 : class, contains x0 (f x0)h:(A -> bool) -> A -> boolHh: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 -> classx:AHx:forall x0 : A, class_of x x0 = h (class_of x) x0Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h gy:AHxy:R x yz:AHxz:c (R x z) = falsefalse = c (R y z)FunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pc:Prop -> boolHc:forall x0 : Prop, c x0 = true <-> x0class_of:=fun a y0 : A => c (R a y0):A -> A -> boolisclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Propclass:={f0 : A -> bool | isclass f0}:Typecontains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Propf:class -> AHf:forall x0 : class, contains x0 (f x0)h:(A -> bool) -> A -> boolHh: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 -> classx:AHx:forall x0 : A, class_of x x0 = h (class_of x) x0Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h gy:AHxy:R x yz:AHxz:c (R x z) = falseHyz:c (R y z) = truefalse = trueFunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pc:Prop -> boolHc:forall x0 : Prop, c x0 = true <-> x0class_of:=fun a y0 : A => c (R a y0):A -> A -> boolisclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Propclass:={f0 : A -> bool | isclass f0}:Typecontains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Propf:class -> AHf:forall x0 : class, contains x0 (f x0)h:(A -> bool) -> A -> boolHh: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 -> classx:AHx:forall x0 : A, class_of x x0 = h (class_of x) x0Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h gy:AHxy:R x yz:AHxz:c (R x z) = falseHyz:c (R y z) = falsefalse = falseFunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pc:Prop -> boolHc:forall x0 : Prop, c x0 = true <-> x0class_of:=fun a y0 : A => c (R a y0):A -> A -> boolisclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Propclass:={f0 : A -> bool | isclass f0}:Typecontains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Propf:class -> AHf:forall x0 : class, contains x0 (f x0)h:(A -> bool) -> A -> boolHh: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 -> classx:AHx:forall x0 : A, class_of x x0 = h (class_of x) x0Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h gy:AHxy:R x yz:AHxz:c (R x z) = falseHyz:c (R y z) = truefalse = trueFunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pc:Prop -> boolHc:forall x0 : Prop, c x0 = true <-> x0class_of:=fun a y0 : A => c (R a y0):A -> A -> boolisclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Propclass:={f0 : A -> bool | isclass f0}:Typecontains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Propf:class -> AHf:forall x0 : class, contains x0 (f x0)h:(A -> bool) -> A -> boolHh: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 -> classx:AHx:forall x0 : A, class_of x x0 = h (class_of x) x0Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h gy:AHxy:R x yz:AHxz:c (R x z) = falseHyz:R y zfalse = trueFunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pc:Prop -> boolHc:forall x0 : Prop, c x0 = true <-> x0class_of:=fun a y0 : A => c (R a y0):A -> A -> boolisclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Propclass:={f0 : A -> bool | isclass f0}:Typecontains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Propf:class -> AHf:forall x0 : class, contains x0 (f x0)h:(A -> bool) -> A -> boolHh: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 -> classx:AHx:forall x0 : A, class_of x x0 = h (class_of x) x0Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h gy:AHxy:R x yz:AHxz:c (R x z) = falseHyz:R y zc (R x z) = truefirstorder.FunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pc:Prop -> boolHc:forall x0 : Prop, c x0 = true <-> x0class_of:=fun a y0 : A => c (R a y0):A -> A -> boolisclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Propclass:={f0 : A -> bool | isclass f0}:Typecontains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Propf:class -> AHf:forall x0 : class, contains x0 (f x0)h:(A -> bool) -> A -> boolHh: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 -> classx:AHx:forall x0 : A, class_of x x0 = h (class_of x) x0Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h gy:AHxy:R x yz:AHxz:c (R x z) = falseHyz:R y zR x zeasy.FunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pc:Prop -> boolHc:forall x0 : Prop, c x0 = true <-> x0class_of:=fun a y0 : A => c (R a y0):A -> A -> boolisclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Propclass:={f0 : A -> bool | isclass f0}:Typecontains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Propf:class -> AHf:forall x0 : class, contains x0 (f x0)h:(A -> bool) -> A -> boolHh: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 -> classx:AHx:forall x0 : A, class_of x x0 = h (class_of x) x0Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h gy:AHxy:R x yz:AHxz:c (R x z) = falseHyz:c (R y z) = falsefalse = falseFunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pc:Prop -> boolHc:forall x0 : Prop, c x0 = true <-> x0class_of:=fun a y0 : A => c (R a y0):A -> A -> boolisclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Propclass:={f0 : A -> bool | isclass f0}:Typecontains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Propf:class -> AHf:forall x0 : class, contains x0 (f x0)h:(A -> bool) -> A -> boolHh: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 -> classx:AHx:forall x0 : A, class_of x x0 = h (class_of x) x0Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h gy:AHxy:R x yHeq1:h (class_of x) = h (class_of y)f' x = f' yFunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pc:Prop -> boolHc:forall x0 : Prop, c x0 = true <-> x0class_of:=fun a y0 : A => c (R a y0):A -> A -> boolisclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Propclass:={f0 : A -> bool | isclass f0}:Typecontains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Propf:class -> AHf:forall x0 : class, contains x0 (f x0)h:(A -> bool) -> A -> boolHh: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 -> classx:AHx:forall x0 : A, class_of x x0 = h (class_of x) x0Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h gy:AHxy:R x yHeq1:h (class_of x) = h (class_of y)rew [isclass] Heq1 in Hisclass x = Hisclass yFunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pc:Prop -> boolHc:forall x0 : Prop, c x0 = true <-> x0class_of:=fun a y0 : A => c (R a y0):A -> A -> boolisclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Propclass:={f0 : A -> bool | isclass f0}:Typecontains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Propf:class -> AHf:forall x0 : class, contains x0 (f x0)h:(A -> bool) -> A -> boolHh: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 -> classx:AHx:forall x0 : A, class_of x x0 = h (class_of x) x0Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h gy:AHxy:R x yHeq1:h (class_of x) = h (class_of y)Heq2:rew [isclass] Heq1 in Hisclass x = Hisclass yf' x = f' yapply proof_irrelevance_cci, EM.FunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pc:Prop -> boolHc:forall x0 : Prop, c x0 = true <-> x0class_of:=fun a y0 : A => c (R a y0):A -> A -> boolisclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Propclass:={f0 : A -> bool | isclass f0}:Typecontains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Propf:class -> AHf:forall x0 : class, contains x0 (f x0)h:(A -> bool) -> A -> boolHh: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 -> classx:AHx:forall x0 : A, class_of x x0 = h (class_of x) x0Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h gy:AHxy:R x yHeq1:h (class_of x) = h (class_of y)rew [isclass] Heq1 in Hisclass x = Hisclass yFunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pc:Prop -> boolHc:forall x0 : Prop, c x0 = true <-> x0class_of:=fun a y0 : A => c (R a y0):A -> A -> boolisclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Propclass:={f0 : A -> bool | isclass f0}:Typecontains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Propf:class -> AHf:forall x0 : class, contains x0 (f x0)h:(A -> bool) -> A -> boolHh: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 -> classx:AHx:forall x0 : A, class_of x x0 = h (class_of x) x0Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h gy:AHxy:R x yHeq1:h (class_of x) = h (class_of y)Heq2:rew [isclass] Heq1 in Hisclass x = Hisclass yf' x = f' yFunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pc:Prop -> boolHc:forall x0 : Prop, c x0 = true <-> x0class_of:=fun a y0 : A => c (R a y0):A -> A -> boolisclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Propclass:={f0 : A -> bool | isclass f0}:Typecontains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Propf:class -> AHf:forall x0 : class, contains x0 (f x0)h:(A -> bool) -> A -> boolHh: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 -> classx:AHx:forall x0 : A, class_of x x0 = h (class_of x) x0Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h gy:AHxy:R x yHeq1:h (class_of x) = h (class_of y)Heq2:rew [isclass] Heq1 in Hisclass x = Hisclass yexist isclass (h (class_of x)) (Hisclass x) = exist isclass (h (class_of y)) (Hisclass y)FunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pc:Prop -> boolHc:forall x0 : Prop, c x0 = true <-> x0class_of:=fun a y0 : A => c (R a y0):A -> A -> boolisclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Propclass:={f0 : A -> bool | isclass f0}:Typecontains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Propf:class -> AHf:forall x0 : class, contains x0 (f x0)h:(A -> bool) -> A -> boolHh: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 -> classx:AHx:forall x0 : A, class_of x x0 = h (class_of x) x0Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h gy:AHxy:R x yHeq1:h (class_of x) = h (class_of y)Heq2:rew [isclass] Heq1 in Hisclass x = Hisclass yexist isclass (h (class_of x)) (Hisclass x) = exist isclass (h (class_of y)) (rew [isclass] Heq1 in Hisclass x)reflexivity. Qed.FunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive RH:forall P : Prop, exists b : bool, b = true <-> Pc:Prop -> boolHc:forall x0 : Prop, c x0 = true <-> x0class_of:=fun a y0 : A => c (R a y0):A -> A -> boolisclass:=fun f0 : A -> bool => exists x0 : A, f0 x0 = true:(A -> bool) -> Propclass:={f0 : A -> bool | isclass f0}:Typecontains:=fun (c0 : class) (a : A) => proj1_sig c0 a = true:class -> A -> Propf:class -> AHf:forall x0 : class, contains x0 (f x0)h:(A -> bool) -> A -> boolHh: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 -> classx:AHx:forall x0 : A, class_of x x0 = h (class_of x) x0Huniqx:forall g : A -> bool, (forall x0 : A, class_of x x0 = g x0) -> h (class_of x) = h gy:AHxy:R x yHeq1:h (class_of x) = h (class_of y)Heq2:rew [isclass] Heq1 in Hisclass x = Hisclass yexist isclass (h (class_of x)) (Hisclass x) = exist isclass (h (class_of x)) (rew [isclass] eq_refl in Hisclass x)FunctionalChoice /\ ExtensionalFunctionRepresentative /\ ExcludedMiddle <-> SetoidFunctionalChoiceFunctionalChoice /\ ExtensionalFunctionRepresentative /\ ExcludedMiddle <-> SetoidFunctionalChoiceFunctionalChoice /\ ExtensionalFunctionRepresentative /\ ExcludedMiddle -> SetoidFunctionalChoiceSetoidFunctionalChoice -> FunctionalChoice /\ ExtensionalFunctionRepresentative /\ ExcludedMiddleFunctionalChoice /\ ExtensionalFunctionRepresentative /\ ExcludedMiddle -> SetoidFunctionalChoiceFunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleSetoidFunctionalChoiceFunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleFunctionalRelReificationFunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleRepresentativeFunctionalChoiceFunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleFunctionalRelReificationapply fun_choice_imp_functional_rel_reification, FunChoice.FunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleA, B:TypeFunctionalRelReification_on A Bnow apply fun_choice_and_ext_functions_repr_and_excluded_middle_imp_setoid_fun_choice.FunChoice:FunctionalChoiceSetoidFunRepr:ExtensionalFunctionRepresentativeEM:ExcludedMiddleRepresentativeFunctionalChoiceSetoidFunctionalChoice -> FunctionalChoice /\ ExtensionalFunctionRepresentative /\ ExcludedMiddleSetoidFunChoice:SetoidFunctionalChoiceFunctionalChoice /\ ExtensionalFunctionRepresentative /\ ExcludedMiddleSetoidFunChoice:SetoidFunctionalChoiceFunctionalChoiceSetoidFunChoice:SetoidFunctionalChoiceExtensionalFunctionRepresentativeSetoidFunChoice:SetoidFunctionalChoiceExcludedMiddlenow intros A B; apply setoid_fun_choice_imp_fun_choice.SetoidFunChoice:SetoidFunctionalChoiceFunctionalChoiceSetoidFunChoice:SetoidFunctionalChoiceExtensionalFunctionRepresentativenow apply setoid_fun_choice_imp_repr_fun_choice.SetoidFunChoice:SetoidFunctionalChoiceRepresentativeFunctionalChoiceSetoidFunChoice:SetoidFunctionalChoiceExcludedMiddlenow apply setoid_fun_choice_imp_repr_fun_choice. Qed. (**********************************************************************)SetoidFunChoice:SetoidFunctionalChoiceRepresentativeFunctionalChoice
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 -> RepresentativeFunctionalChoiceFunctionalChoice -> ExtensionalPredicateRepresentative -> ProofIrrelevance -> RepresentativeFunctionalChoiceFunChoice:FunctionalChoicePredExtRepr:ExtensionalPredicateRepresentativePI:ProofIrrelevanceA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive Rexists f : A -> A, forall x : A, R x (f x) /\ (forall x' : A, R x x' -> f x = f x')FunChoice:FunctionalChoicePredExtRepr:ExtensionalPredicateRepresentativePI:ProofIrrelevanceA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive Risclass:=fun P : A -> Prop => exists x : A, P x:(A -> Prop) -> Propexists f : A -> A, forall x : A, R x (f x) /\ (forall x' : A, R x x' -> f x = f x')FunChoice:FunctionalChoicePredExtRepr:ExtensionalPredicateRepresentativePI:ProofIrrelevanceA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive Risclass:=fun P : A -> Prop => exists x : A, P x:(A -> Prop) -> Propclass:={P : A -> Prop | isclass P}:Typeexists f : A -> A, forall x : A, R x (f x) /\ (forall x' : A, R x x' -> f x = f x')FunChoice:FunctionalChoicePredExtRepr:ExtensionalPredicateRepresentativePI:ProofIrrelevanceA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive Risclass:=fun P : A -> Prop => exists x : A, P x:(A -> Prop) -> Propclass:={P : A -> Prop | isclass P}:Typecontains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Propexists f : A -> A, forall x : A, R x (f x) /\ (forall x' : A, R x x' -> f x = f x')FunChoice:FunctionalChoicePredExtRepr:ExtensionalPredicateRepresentativePI:ProofIrrelevanceA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive Risclass:=fun P : A -> Prop => exists x : A, P x:(A -> Prop) -> Propclass:={P : A -> Prop | isclass P}:Typecontains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Propclass_of:=fun a : A => R a:A -> A -> Propexists f : A -> A, forall x : A, R x (f x) /\ (forall x' : A, R x x' -> f x = f x')FunChoice:FunctionalChoicePredExtRepr:ExtensionalPredicateRepresentativePI:ProofIrrelevanceA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive Risclass:=fun P : A -> Prop => exists x : A, P x:(A -> Prop) -> Propclass:={P : A -> Prop | isclass P}:Typecontains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Propclass_of:=fun a : A => R a:A -> A -> Propforall x : class, exists y : A, contains x yFunChoice:FunctionalChoicePredExtRepr:ExtensionalPredicateRepresentativePI:ProofIrrelevanceA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive Risclass:=fun P : A -> Prop => exists x : A, P x:(A -> Prop) -> Propclass:={P : A -> Prop | isclass P}:Typecontains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Propclass_of:=fun a : A => R a:A -> A -> Propf:class -> AHf: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:FunctionalChoicePredExtRepr:ExtensionalPredicateRepresentativePI:ProofIrrelevanceA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive Risclass:=fun P : A -> Prop => exists x : A, P x:(A -> Prop) -> Propclass:={P : A -> Prop | isclass P}:Typecontains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Propclass_of:=fun a : A => R a:A -> A -> Propforall x : class, exists y : A, contains x yapply proj2_sig.FunChoice:FunctionalChoicePredExtRepr:ExtensionalPredicateRepresentativePI:ProofIrrelevanceA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive Risclass:=fun P : A -> Prop => exists x : A, P x:(A -> Prop) -> Propclass:={P : A -> Prop | isclass P}:Typecontains:=fun (c0 : class) (a : A) => proj1_sig c0 a:class -> A -> Propclass_of:=fun a : A => R a:A -> A -> Propc:classexists y : A, contains c yFunChoice:FunctionalChoicePredExtRepr:ExtensionalPredicateRepresentativePI:ProofIrrelevanceA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive Risclass:=fun P : A -> Prop => exists x : A, P x:(A -> Prop) -> Propclass:={P : A -> Prop | isclass P}:Typecontains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Propclass_of:=fun a : A => R a:A -> A -> Propf:class -> AHf: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:FunctionalChoicePredExtRepr:ExtensionalPredicateRepresentativePI:ProofIrrelevanceA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive Risclass:=fun P : A -> Prop => exists x : A, P x:(A -> Prop) -> Propclass:={P : A -> Prop | isclass P}:Typecontains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Propclass_of:=fun a : A => R a:A -> A -> Propf:class -> AHf:forall x : class, contains x (f x)h:(A -> Prop) -> A -> PropHh: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:FunctionalChoicePredExtRepr:ExtensionalPredicateRepresentativePI:ProofIrrelevanceA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive Risclass:=fun P : A -> Prop => exists x : A, P x:(A -> Prop) -> Propclass:={P : A -> Prop | isclass P}:Typecontains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Propclass_of:=fun a : A => R a:A -> A -> Propf:class -> AHf:forall x : class, contains x (f x)h:(A -> Prop) -> A -> PropHh: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:FunctionalChoicePredExtRepr:ExtensionalPredicateRepresentativePI:ProofIrrelevanceA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive Risclass:=fun P : A -> Prop => exists x : A, P x:(A -> Prop) -> Propclass:={P : A -> Prop | isclass P}:Typecontains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Propclass_of:=fun a : A => R a:A -> A -> Propf:class -> AHf:forall x : class, contains x (f x)h:(A -> Prop) -> A -> PropHh: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:FunctionalChoicePredExtRepr:ExtensionalPredicateRepresentativePI:ProofIrrelevanceA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive Risclass:=fun P : A -> Prop => exists x : A, P x:(A -> Prop) -> Propclass:={P : A -> Prop | isclass P}:Typecontains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Propclass_of:=fun a : A => R a:A -> A -> Propf:class -> AHf:forall x : class, contains x (f x)h:(A -> Prop) -> A -> PropHh: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:FunctionalChoicePredExtRepr:ExtensionalPredicateRepresentativePI:ProofIrrelevanceA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive Risclass:=fun P : A -> Prop => exists x : A, P x:(A -> Prop) -> Propclass:={P : A -> Prop | isclass P}:Typecontains:=fun (c : class) (a0 : A) => proj1_sig c a0:class -> A -> Propclass_of:=fun a0 : A => R a0:A -> A -> Propf:class -> AHf:forall x : class, contains x (f x)h:(A -> Prop) -> A -> PropHh: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:Aisclass (h (class_of a))FunChoice:FunctionalChoicePredExtRepr:ExtensionalPredicateRepresentativePI:ProofIrrelevanceA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive Risclass:=fun P : A -> Prop => exists x : A, P x:(A -> Prop) -> Propclass:={P : A -> Prop | isclass P}:Typecontains:=fun (c : class) (a0 : A) => proj1_sig c a0:class -> A -> Propclass_of:=fun a0 : A => R a0:A -> A -> Propf:class -> AHf:forall x : class, contains x (f x)h:(A -> Prop) -> A -> PropHh: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:Ah (class_of a) arewrite <- Ha; apply Hrefl.FunChoice:FunctionalChoicePredExtRepr:ExtensionalPredicateRepresentativePI:ProofIrrelevanceA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive Risclass:=fun P : A -> Prop => exists x : A, P x:(A -> Prop) -> Propclass:={P : A -> Prop | isclass P}:Typecontains:=fun (c : class) (a0 : A) => proj1_sig c a0:class -> A -> Propclass_of:=fun a0 : A => R a0:A -> A -> Propf:class -> AHf:forall x : class, contains x (f x)h:(A -> Prop) -> A -> PropHh: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:AHa:forall x : A, class_of a x <-> h (class_of a) xHuniqa:forall Q : A -> Prop, (forall x : A, class_of a x <-> Q x) -> h (class_of a) = h Qh (class_of a) aFunChoice:FunctionalChoicePredExtRepr:ExtensionalPredicateRepresentativePI:ProofIrrelevanceA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive Risclass:=fun P : A -> Prop => exists x : A, P x:(A -> Prop) -> Propclass:={P : A -> Prop | isclass P}:Typecontains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Propclass_of:=fun a : A => R a:A -> A -> Propf:class -> AHf:forall x : class, contains x (f x)h:(A -> Prop) -> A -> PropHh: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:FunctionalChoicePredExtRepr:ExtensionalPredicateRepresentativePI:ProofIrrelevanceA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive Risclass:=fun P : A -> Prop => exists x : A, P x:(A -> Prop) -> Propclass:={P : A -> Prop | isclass P}:Typecontains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Propclass_of:=fun a : A => R a:A -> A -> Propf:class -> AHf:forall x : class, contains x (f x)h:(A -> Prop) -> A -> PropHh: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 -> classexists f0 : A -> A, forall x : A, R x (f0 x) /\ (forall x' : A, R x x' -> f0 x = f0 x')FunChoice:FunctionalChoicePredExtRepr:ExtensionalPredicateRepresentativePI:ProofIrrelevanceA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive Risclass:=fun P : A -> Prop => exists x : A, P x:(A -> Prop) -> Propclass:={P : A -> Prop | isclass P}:Typecontains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Propclass_of:=fun a : A => R a:A -> A -> Propf:class -> AHf:forall x : class, contains x (f x)h:(A -> Prop) -> A -> PropHh: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 -> classforall x : A, R x (f (f' x)) /\ (forall x' : A, R x x' -> f (f' x) = f (f' x'))FunChoice:FunctionalChoicePredExtRepr:ExtensionalPredicateRepresentativePI:ProofIrrelevanceA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive Risclass:=fun P : A -> Prop => exists x0 : A, P x0:(A -> Prop) -> Propclass:={P : A -> Prop | isclass P}:Typecontains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Propclass_of:=fun a : A => R a:A -> A -> Propf:class -> AHf:forall x0 : class, contains x0 (f x0)h:(A -> Prop) -> A -> PropHh: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 -> classx:AR x (f (f' x)) /\ (forall x' : A, R x x' -> f (f' x) = f (f' x'))FunChoice:FunctionalChoicePredExtRepr:ExtensionalPredicateRepresentativePI:ProofIrrelevanceA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive Risclass:=fun P : A -> Prop => exists x0 : A, P x0:(A -> Prop) -> Propclass:={P : A -> Prop | isclass P}:Typecontains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Propclass_of:=fun a : A => R a:A -> A -> Propf:class -> AHf:forall x0 : class, contains x0 (f x0)h:(A -> Prop) -> A -> PropHh: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 -> classx:AHx:forall x0 : A, class_of x x0 <-> h (class_of x) x0Huniqx:forall Q : A -> Prop, (forall x0 : A, class_of x x0 <-> Q x0) -> h (class_of x) = h QR x (f (f' x)) /\ (forall x' : A, R x x' -> f (f' x) = f (f' x'))FunChoice:FunctionalChoicePredExtRepr:ExtensionalPredicateRepresentativePI:ProofIrrelevanceA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive Risclass:=fun P : A -> Prop => exists x0 : A, P x0:(A -> Prop) -> Propclass:={P : A -> Prop | isclass P}:Typecontains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Propclass_of:=fun a : A => R a:A -> A -> Propf:class -> AHf:forall x0 : class, contains x0 (f x0)h:(A -> Prop) -> A -> PropHh: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 -> classx:AHx:forall x0 : A, class_of x x0 <-> h (class_of x) x0Huniqx:forall Q : A -> Prop, (forall x0 : A, class_of x x0 <-> Q x0) -> h (class_of x) = h QR x (f (f' x))FunChoice:FunctionalChoicePredExtRepr:ExtensionalPredicateRepresentativePI:ProofIrrelevanceA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive Risclass:=fun P : A -> Prop => exists x0 : A, P x0:(A -> Prop) -> Propclass:={P : A -> Prop | isclass P}:Typecontains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Propclass_of:=fun a : A => R a:A -> A -> Propf:class -> AHf:forall x0 : class, contains x0 (f x0)h:(A -> Prop) -> A -> PropHh: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 -> classx:AHx:forall x0 : A, class_of x x0 <-> h (class_of x) x0Huniqx:forall Q : A -> Prop, (forall x0 : A, class_of x x0 <-> Q x0) -> h (class_of x) = h Qforall x' : A, R x x' -> f (f' x) = f (f' x')FunChoice:FunctionalChoicePredExtRepr:ExtensionalPredicateRepresentativePI:ProofIrrelevanceA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive Risclass:=fun P : A -> Prop => exists x0 : A, P x0:(A -> Prop) -> Propclass:={P : A -> Prop | isclass P}:Typecontains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Propclass_of:=fun a : A => R a:A -> A -> Propf:class -> AHf:forall x0 : class, contains x0 (f x0)h:(A -> Prop) -> A -> PropHh: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 -> classx:AHx:forall x0 : A, class_of x x0 <-> h (class_of x) x0Huniqx:forall Q : A -> Prop, (forall x0 : A, class_of x x0 <-> Q x0) -> h (class_of x) = h QR x (f (f' x))FunChoice:FunctionalChoicePredExtRepr:ExtensionalPredicateRepresentativePI:ProofIrrelevanceA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive Risclass:=fun P : A -> Prop => exists x0 : A, P x0:(A -> Prop) -> Propclass:={P : A -> Prop | isclass P}:Typecontains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Propclass_of:=fun a : A => R a:A -> A -> Propf:class -> Ah:(A -> Prop) -> A -> PropHisclass:forall a : A, isclass (h (class_of a))f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> classx:AHf: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) x0Huniqx:forall Q : A -> Prop, (forall x0 : A, class_of x x0 <-> Q x0) -> h (class_of x) = h QR x (f (f' x))FunChoice:FunctionalChoicePredExtRepr:ExtensionalPredicateRepresentativePI:ProofIrrelevanceA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive Risclass:=fun P : A -> Prop => exists x0 : A, P x0:(A -> Prop) -> Propclass:={P : A -> Prop | isclass P}:Typecontains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Propclass_of:=fun a : A => R a:A -> A -> Propf:class -> Ah:(A -> Prop) -> A -> PropHisclass:forall a : A, isclass (h (class_of a))f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> classx:AHf: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) x0Huniqx:forall Q : A -> Prop, (forall x0 : A, class_of x x0 <-> Q x0) -> h (class_of x) = h QR x (f (f' x))assumption.FunChoice:FunctionalChoicePredExtRepr:ExtensionalPredicateRepresentativePI:ProofIrrelevanceA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive Risclass:=fun P : A -> Prop => exists x0 : A, P x0:(A -> Prop) -> Propclass:={P : A -> Prop | isclass P}:Typecontains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Propclass_of:=fun a : A => R a:A -> A -> Propf:class -> Ah:(A -> Prop) -> A -> PropHisclass:forall a : A, isclass (h (class_of a))f':=fun a : A => exist isclass (h (class_of a)) (Hisclass a) : class:A -> classx:AHf: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) x0Huniqx:forall Q : A -> Prop, (forall x0 : A, class_of x x0 <-> Q x0) -> h (class_of x) = h QR x (f (f' x))FunChoice:FunctionalChoicePredExtRepr:ExtensionalPredicateRepresentativePI:ProofIrrelevanceA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive Risclass:=fun P : A -> Prop => exists x0 : A, P x0:(A -> Prop) -> Propclass:={P : A -> Prop | isclass P}:Typecontains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Propclass_of:=fun a : A => R a:A -> A -> Propf:class -> AHf:forall x0 : class, contains x0 (f x0)h:(A -> Prop) -> A -> PropHh: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 -> classx:AHx:forall x0 : A, class_of x x0 <-> h (class_of x) x0Huniqx:forall Q : A -> Prop, (forall x0 : A, class_of x x0 <-> Q x0) -> h (class_of x) = h Qforall x' : A, R x x' -> f (f' x) = f (f' x')FunChoice:FunctionalChoicePredExtRepr:ExtensionalPredicateRepresentativePI:ProofIrrelevanceA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive Risclass:=fun P : A -> Prop => exists x0 : A, P x0:(A -> Prop) -> Propclass:={P : A -> Prop | isclass P}:Typecontains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Propclass_of:=fun a : A => R a:A -> A -> Propf:class -> AHf:forall x0 : class, contains x0 (f x0)h:(A -> Prop) -> A -> PropHh: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 -> classx:AHx:forall x0 : A, class_of x x0 <-> h (class_of x) x0Huniqx:forall Q : A -> Prop, (forall x0 : A, class_of x x0 <-> Q x0) -> h (class_of x) = h Qy:AHxy:R x yf (f' x) = f (f' y)FunChoice:FunctionalChoicePredExtRepr:ExtensionalPredicateRepresentativePI:ProofIrrelevanceA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive Risclass:=fun P : A -> Prop => exists x0 : A, P x0:(A -> Prop) -> Propclass:={P : A -> Prop | isclass P}:Typecontains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Propclass_of:=fun a : A => R a:A -> A -> Propf:class -> AHf:forall x0 : class, contains x0 (f x0)h:(A -> Prop) -> A -> PropHh: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 -> classx:AHx:forall x0 : A, class_of x x0 <-> h (class_of x) x0Huniqx:forall Q : A -> Prop, (forall x0 : A, class_of x x0 <-> Q x0) -> h (class_of x) = h Qy:AHxy:R x yf' x = f' yFunChoice:FunctionalChoicePredExtRepr:ExtensionalPredicateRepresentativePI:ProofIrrelevanceA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive Risclass:=fun P : A -> Prop => exists x0 : A, P x0:(A -> Prop) -> Propclass:={P : A -> Prop | isclass P}:Typecontains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Propclass_of:=fun a : A => R a:A -> A -> Propf:class -> AHf:forall x0 : class, contains x0 (f x0)h:(A -> Prop) -> A -> PropHh: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 -> classx:AHx:forall x0 : A, class_of x x0 <-> h (class_of x) x0Huniqx:forall Q : A -> Prop, (forall x0 : A, class_of x x0 <-> Q x0) -> h (class_of x) = h Qy:AHxy:R x yh (class_of x) = h (class_of y)FunChoice:FunctionalChoicePredExtRepr:ExtensionalPredicateRepresentativePI:ProofIrrelevanceA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive Risclass:=fun P : A -> Prop => exists x0 : A, P x0:(A -> Prop) -> Propclass:={P : A -> Prop | isclass P}:Typecontains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Propclass_of:=fun a : A => R a:A -> A -> Propf:class -> AHf:forall x0 : class, contains x0 (f x0)h:(A -> Prop) -> A -> PropHh: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 -> classx:AHx:forall x0 : A, class_of x x0 <-> h (class_of x) x0Huniqx:forall Q : A -> Prop, (forall x0 : A, class_of x x0 <-> Q x0) -> h (class_of x) = h Qy:AHxy:R x yHeq1:h (class_of x) = h (class_of y)f' x = f' yFunChoice:FunctionalChoicePredExtRepr:ExtensionalPredicateRepresentativePI:ProofIrrelevanceA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive Risclass:=fun P : A -> Prop => exists x0 : A, P x0:(A -> Prop) -> Propclass:={P : A -> Prop | isclass P}:Typecontains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Propclass_of:=fun a : A => R a:A -> A -> Propf:class -> AHf:forall x0 : class, contains x0 (f x0)h:(A -> Prop) -> A -> PropHh: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 -> classx:AHx:forall x0 : A, class_of x x0 <-> h (class_of x) x0Huniqx:forall Q : A -> Prop, (forall x0 : A, class_of x x0 <-> Q x0) -> h (class_of x) = h Qy:AHxy:R x yh (class_of x) = h (class_of y)FunChoice:FunctionalChoicePredExtRepr:ExtensionalPredicateRepresentativePI:ProofIrrelevanceA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive Risclass:=fun P : A -> Prop => exists x0 : A, P x0:(A -> Prop) -> Propclass:={P : A -> Prop | isclass P}:Typecontains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Propclass_of:=fun a : A => R a:A -> A -> Propf:class -> AHf:forall x0 : class, contains x0 (f x0)h:(A -> Prop) -> A -> PropHh: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 -> classx:AHx:forall x0 : A, class_of x x0 <-> h (class_of x) x0Huniqx:forall Q : A -> Prop, (forall x0 : A, class_of x x0 <-> Q x0) -> h (class_of x) = h Qy:AHxy:R x yforall x0 : A, class_of x x0 <-> class_of y x0FunChoice:FunctionalChoicePredExtRepr:ExtensionalPredicateRepresentativePI:ProofIrrelevanceA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive Risclass:=fun P : A -> Prop => exists x0 : A, P x0:(A -> Prop) -> Propclass:={P : A -> Prop | isclass P}:Typecontains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Propclass_of:=fun a : A => R a:A -> A -> Propf:class -> AHf:forall x0 : class, contains x0 (f x0)h:(A -> Prop) -> A -> PropHh: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 -> classx:AHx:forall x0 : A, class_of x x0 <-> h (class_of x) x0Huniqx:forall Q : A -> Prop, (forall x0 : A, class_of x x0 <-> Q x0) -> h (class_of x) = h Qy:AHxy:R x yz:Aclass_of x z <-> class_of y zfirstorder.FunChoice:FunctionalChoicePredExtRepr:ExtensionalPredicateRepresentativePI:ProofIrrelevanceA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive Risclass:=fun P : A -> Prop => exists x0 : A, P x0:(A -> Prop) -> Propclass:={P : A -> Prop | isclass P}:Typecontains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Propclass_of:=fun a : A => R a:A -> A -> Propf:class -> AHf:forall x0 : class, contains x0 (f x0)h:(A -> Prop) -> A -> PropHh: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 -> classx:AHx:forall x0 : A, class_of x x0 <-> h (class_of x) x0Huniqx:forall Q : A -> Prop, (forall x0 : A, class_of x x0 <-> Q x0) -> h (class_of x) = h Qy:AHxy:R x yz:AR x z <-> R y zFunChoice:FunctionalChoicePredExtRepr:ExtensionalPredicateRepresentativePI:ProofIrrelevanceA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive Risclass:=fun P : A -> Prop => exists x0 : A, P x0:(A -> Prop) -> Propclass:={P : A -> Prop | isclass P}:Typecontains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Propclass_of:=fun a : A => R a:A -> A -> Propf:class -> AHf:forall x0 : class, contains x0 (f x0)h:(A -> Prop) -> A -> PropHh: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 -> classx:AHx:forall x0 : A, class_of x x0 <-> h (class_of x) x0Huniqx:forall Q : A -> Prop, (forall x0 : A, class_of x x0 <-> Q x0) -> h (class_of x) = h Qy:AHxy:R x yHeq1:h (class_of x) = h (class_of y)f' x = f' yFunChoice:FunctionalChoicePredExtRepr:ExtensionalPredicateRepresentativePI:ProofIrrelevanceA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive Risclass:=fun P : A -> Prop => exists x0 : A, P x0:(A -> Prop) -> Propclass:={P : A -> Prop | isclass P}:Typecontains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Propclass_of:=fun a : A => R a:A -> A -> Propf:class -> AHf:forall x0 : class, contains x0 (f x0)h:(A -> Prop) -> A -> PropHh: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 -> classx:AHx:forall x0 : A, class_of x x0 <-> h (class_of x) x0Huniqx:forall Q : A -> Prop, (forall x0 : A, class_of x x0 <-> Q x0) -> h (class_of x) = h Qy:AHxy:R x yHeq1:h (class_of x) = h (class_of y)rew [isclass] Heq1 in Hisclass x = Hisclass yFunChoice:FunctionalChoicePredExtRepr:ExtensionalPredicateRepresentativePI:ProofIrrelevanceA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive Risclass:=fun P : A -> Prop => exists x0 : A, P x0:(A -> Prop) -> Propclass:={P : A -> Prop | isclass P}:Typecontains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Propclass_of:=fun a : A => R a:A -> A -> Propf:class -> AHf:forall x0 : class, contains x0 (f x0)h:(A -> Prop) -> A -> PropHh: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 -> classx:AHx:forall x0 : A, class_of x x0 <-> h (class_of x) x0Huniqx:forall Q : A -> Prop, (forall x0 : A, class_of x x0 <-> Q x0) -> h (class_of x) = h Qy:AHxy:R x yHeq1:h (class_of x) = h (class_of y)Heq2:rew [isclass] Heq1 in Hisclass x = Hisclass yf' x = f' yapply PI.FunChoice:FunctionalChoicePredExtRepr:ExtensionalPredicateRepresentativePI:ProofIrrelevanceA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive Risclass:=fun P : A -> Prop => exists x0 : A, P x0:(A -> Prop) -> Propclass:={P : A -> Prop | isclass P}:Typecontains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Propclass_of:=fun a : A => R a:A -> A -> Propf:class -> AHf:forall x0 : class, contains x0 (f x0)h:(A -> Prop) -> A -> PropHh: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 -> classx:AHx:forall x0 : A, class_of x x0 <-> h (class_of x) x0Huniqx:forall Q : A -> Prop, (forall x0 : A, class_of x x0 <-> Q x0) -> h (class_of x) = h Qy:AHxy:R x yHeq1:h (class_of x) = h (class_of y)rew [isclass] Heq1 in Hisclass x = Hisclass yFunChoice:FunctionalChoicePredExtRepr:ExtensionalPredicateRepresentativePI:ProofIrrelevanceA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive Risclass:=fun P : A -> Prop => exists x0 : A, P x0:(A -> Prop) -> Propclass:={P : A -> Prop | isclass P}:Typecontains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Propclass_of:=fun a : A => R a:A -> A -> Propf:class -> AHf:forall x0 : class, contains x0 (f x0)h:(A -> Prop) -> A -> PropHh: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 -> classx:AHx:forall x0 : A, class_of x x0 <-> h (class_of x) x0Huniqx:forall Q : A -> Prop, (forall x0 : A, class_of x x0 <-> Q x0) -> h (class_of x) = h Qy:AHxy:R x yHeq1:h (class_of x) = h (class_of y)Heq2:rew [isclass] Heq1 in Hisclass x = Hisclass yf' x = f' yFunChoice:FunctionalChoicePredExtRepr:ExtensionalPredicateRepresentativePI:ProofIrrelevanceA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive Risclass:=fun P : A -> Prop => exists x0 : A, P x0:(A -> Prop) -> Propclass:={P : A -> Prop | isclass P}:Typecontains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Propclass_of:=fun a : A => R a:A -> A -> Propf:class -> AHf:forall x0 : class, contains x0 (f x0)h:(A -> Prop) -> A -> PropHh: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 -> classx:AHx:forall x0 : A, class_of x x0 <-> h (class_of x) x0Huniqx:forall Q : A -> Prop, (forall x0 : A, class_of x x0 <-> Q x0) -> h (class_of x) = h Qy:AHxy:R x yHeq1:h (class_of x) = h (class_of y)Heq2:rew [isclass] Heq1 in Hisclass x = Hisclass yexist isclass (h (class_of x)) (Hisclass x) = exist isclass (h (class_of y)) (Hisclass y)FunChoice:FunctionalChoicePredExtRepr:ExtensionalPredicateRepresentativePI:ProofIrrelevanceA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive Risclass:=fun P : A -> Prop => exists x0 : A, P x0:(A -> Prop) -> Propclass:={P : A -> Prop | isclass P}:Typecontains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Propclass_of:=fun a : A => R a:A -> A -> Propf:class -> AHf:forall x0 : class, contains x0 (f x0)h:(A -> Prop) -> A -> PropHh: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 -> classx:AHx:forall x0 : A, class_of x x0 <-> h (class_of x) x0Huniqx:forall Q : A -> Prop, (forall x0 : A, class_of x x0 <-> Q x0) -> h (class_of x) = h Qy:AHxy:R x yHeq1:h (class_of x) = h (class_of y)Heq2:rew [isclass] Heq1 in Hisclass x = Hisclass yexist isclass (h (class_of x)) (Hisclass x) = exist isclass (h (class_of y)) (rew [isclass] Heq1 in Hisclass x)reflexivity. Qed.FunChoice:FunctionalChoicePredExtRepr:ExtensionalPredicateRepresentativePI:ProofIrrelevanceA:TypeR:A -> A -> PropHrefl:Reflexive RHsym:Symmetric RHtrans:Transitive Risclass:=fun P : A -> Prop => exists x0 : A, P x0:(A -> Prop) -> Propclass:={P : A -> Prop | isclass P}:Typecontains:=fun (c : class) (a : A) => proj1_sig c a:class -> A -> Propclass_of:=fun a : A => R a:A -> A -> Propf:class -> AHf:forall x0 : class, contains x0 (f x0)h:(A -> Prop) -> A -> PropHh: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 -> classx:AHx:forall x0 : A, class_of x x0 <-> h (class_of x) x0Huniqx:forall Q : A -> Prop, (forall x0 : A, class_of x x0 <-> Q x0) -> h (class_of x) = h Qy:AHxy:R x yHeq1:h (class_of x) = h (class_of y)Heq2:rew [isclass] Heq1 in Hisclass x = Hisclass yexist isclass (h (class_of x)) (Hisclass x) = exist isclass (h (class_of x)) (rew [isclass] eq_refl in Hisclass x)FunctionalChoice /\ ExtensionalPredicateRepresentative /\ ProofIrrelevance <-> SetoidFunctionalChoiceFunctionalChoice /\ ExtensionalPredicateRepresentative /\ ProofIrrelevance <-> SetoidFunctionalChoiceFunctionalChoice /\ ExtensionalPredicateRepresentative /\ ProofIrrelevance -> SetoidFunctionalChoiceSetoidFunctionalChoice -> FunctionalChoice /\ ExtensionalPredicateRepresentative /\ ProofIrrelevanceFunctionalChoice /\ ExtensionalPredicateRepresentative /\ ProofIrrelevance -> SetoidFunctionalChoiceFunChoice:FunctionalChoiceExtPredRepr:ExtensionalPredicateRepresentativePI:ProofIrrelevanceSetoidFunctionalChoiceFunChoice:FunctionalChoiceExtPredRepr:ExtensionalPredicateRepresentativePI:ProofIrrelevanceFunctionalRelReificationFunChoice:FunctionalChoiceExtPredRepr:ExtensionalPredicateRepresentativePI:ProofIrrelevanceRepresentativeFunctionalChoiceFunChoice:FunctionalChoiceExtPredRepr:ExtensionalPredicateRepresentativePI:ProofIrrelevanceFunctionalRelReificationnow apply fun_choice_imp_functional_rel_reification.FunChoice:FunctionalChoiceExtPredRepr:ExtensionalPredicateRepresentativePI:ProofIrrelevanceA, B:TypeFunctionalRelReification_on A Bnow apply fun_choice_and_ext_pred_ext_and_proof_irrel_imp_setoid_fun_choice.FunChoice:FunctionalChoiceExtPredRepr:ExtensionalPredicateRepresentativePI:ProofIrrelevanceRepresentativeFunctionalChoiceSetoidFunctionalChoice -> FunctionalChoice /\ ExtensionalPredicateRepresentative /\ ProofIrrelevanceSetoidFunChoice:SetoidFunctionalChoiceFunctionalChoice /\ ExtensionalPredicateRepresentative /\ ProofIrrelevanceSetoidFunChoice:SetoidFunctionalChoiceFunctionalChoiceSetoidFunChoice:SetoidFunctionalChoiceExtensionalPredicateRepresentativeSetoidFunChoice:SetoidFunctionalChoiceProofIrrelevancenow intros A B; apply setoid_fun_choice_imp_fun_choice.SetoidFunChoice:SetoidFunctionalChoiceFunctionalChoiceSetoidFunChoice:SetoidFunctionalChoiceExtensionalPredicateRepresentativenow apply setoid_fun_choice_imp_repr_fun_choice.SetoidFunChoice:SetoidFunctionalChoiceRepresentativeFunctionalChoiceSetoidFunChoice:SetoidFunctionalChoiceProofIrrelevanceSetoidFunChoice:SetoidFunctionalChoiceforall (A : Prop) (a1 a2 : A), a1 = a2SetoidFunChoice:SetoidFunctionalChoiceforall A : Prop, A \/ ~ Anow apply setoid_fun_choice_imp_repr_fun_choice. Qed. (**********************************************************************)SetoidFunChoice:SetoidFunctionalChoiceRepresentativeFunctionalChoice
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).