From 9c21392c7957a0a1a66e5269022d5991649a38b5 Mon Sep 17 00:00:00 2001 From: Hugo Herbelin Date: Sun, 2 Oct 2016 22:02:52 +0200 Subject: Adding a file providing extensional choice (i.e. choice over setoids). MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Also integrating suggestions from Théo. --- theories/Logic/ExtensionalFunctionRepresentative.v | 24 +++++++++ theories/Logic/SetoidChoice.v | 60 ++++++++++++++++++++++ theories/Logic/vo.itarget | 4 +- 3 files changed, 87 insertions(+), 1 deletion(-) create mode 100644 theories/Logic/ExtensionalFunctionRepresentative.v create mode 100644 theories/Logic/SetoidChoice.v (limited to 'theories/Logic') diff --git a/theories/Logic/ExtensionalFunctionRepresentative.v b/theories/Logic/ExtensionalFunctionRepresentative.v new file mode 100644 index 000000000..a9da68e16 --- /dev/null +++ b/theories/Logic/ExtensionalFunctionRepresentative.v @@ -0,0 +1,24 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* B), + (forall x, f x = repr f x) /\ + (forall g, (forall x, f x = g x) -> repr f = repr g). diff --git a/theories/Logic/SetoidChoice.v b/theories/Logic/SetoidChoice.v new file mode 100644 index 000000000..84432dda1 --- /dev/null +++ b/theories/Logic/SetoidChoice.v @@ -0,0 +1,60 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* 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'). +Proof. + apply setoid_functional_choice_first_characterization. split; [|split]. + - exact choice. + - exact extensional_function_representative. + - exact classic. +Qed. + +Theorem representative_choice : + forall A (R:A->A->Prop), (Equivalence R) -> + exists f : A->A, forall x : A, R x (f x) /\ forall x', R x x' -> f x = f x'. +Proof. + apply setoid_fun_choice_imp_repr_fun_choice. + exact setoid_choice. +Qed. diff --git a/theories/Logic/vo.itarget b/theories/Logic/vo.itarget index 323597395..71d01c1fb 100644 --- a/theories/Logic/vo.itarget +++ b/theories/Logic/vo.itarget @@ -20,6 +20,7 @@ WeakFan.vo WKL.vo FunctionalExtensionality.vo ExtensionalityFacts.vo +ExtensionalFunctionRepresentative.vo Hurkens.vo IndefiniteDescription.vo JMeq.vo @@ -27,4 +28,5 @@ ProofIrrelevanceFacts.vo ProofIrrelevance.vo RelationalChoice.vo SetIsType.vo -FinFun.vo \ No newline at end of file +SetoidChoice.vo +FinFun.vo -- cgit v1.2.3