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 about extensionality
We investigate the relations between the following extensionality principles
Table of contents
1. Definitions
2. Functional extensionality <-> Equality of projections from diagonal
3. Functional extensionality <-> Unicity of inverse bijections
4. Functional extensionality <-> Bijectivity of bijective composition
Set Implicit Arguments.

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

Definitions

Being an inverse
Definition is_inverse A B f g := (forall a:A, g (f a) = a) /\ (forall b:B, f (g b) = b).
The diagonal over A and the one-one correspondence with A
#[universes(template)]
Record Delta A := { pi1:A; pi2:A; eq:pi1=pi2 }.

Definition delta {A} (a:A) := {|pi1 := a; pi2 := a; eq := eq_refl a |}.

Arguments pi1 {A} _.
Arguments pi2 {A} _.


forall (A : Type) (x : Delta A), pi1 x = pi2 x

forall (A : Type) (x : Delta A), pi1 x = pi2 x
destruct x as (a1,a2,Heq); assumption. Qed.

forall A : Type, is_inverse delta pi1

forall A : Type, is_inverse delta pi1
split; [trivial|]; destruct b as (a1,a2,[]); reflexivity. Qed.

forall A : Type, is_inverse delta pi2

forall A : Type, is_inverse delta pi2
split; [trivial|]; destruct b as (a1,a2,[]); reflexivity. Qed.
Functional extensionality
Notation FunctionalExtensionality := 
  (forall A B (f g : A -> B), (forall x, f x = g x) -> f = g).
Equality of projections from diagonal
Notation EqDeltaProjs := (forall A, pi1 = pi2 :> (Delta A -> A)).
Unicity of bijection inverse
Notation UniqueInverse := (forall A B (f:A->B) g1 g2, is_inverse f g1 -> is_inverse f g2 -> g1 = g2).
Bijectivity of bijective composition
Definition action A B C (f:A->B) := (fun h:B->C => fun x => h (f x)).

Notation BijectivityBijectiveComp := (forall A B C (f:A->B) g,
  is_inverse f g -> is_inverse (A:=B->C) (action f) (action g)).

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

Functional extensionality <-> Equality of projections from diagonal


FunctionalExtensionality <-> EqDeltaProjs

FunctionalExtensionality <-> EqDeltaProjs

FunctionalExtensionality -> EqDeltaProjs

EqDeltaProjs -> FunctionalExtensionality

FunctionalExtensionality -> EqDeltaProjs
intros FunExt *; apply FunExt, diagonal_projs_same_behavior.

EqDeltaProjs -> FunctionalExtensionality
EqProjs:EqDeltaProjs
A, B:Type
f, g:A -> B
H:forall x : A, f x = g x

(fun x : A => pi1 {| pi1 := f x; pi2 := g x; eq := H x |}) = g
rewrite EqProjs; reflexivity. Qed. (**********************************************************************)

Functional extensionality <-> Unicity of bijection inverse


FunctionalExtensionality -> UniqueInverse

FunctionalExtensionality -> UniqueInverse
FunExt:FunctionalExtensionality
A, B:Type
f:A -> B
g1, g2:B -> A
Hg1f:forall a : A, g1 (f a) = a
Hfg1:forall b : B, f (g1 b) = b
Hg2f:forall a : A, g2 (f a) = a
Hfg2:forall b : B, f (g2 b) = b

g1 = g2
FunExt:FunctionalExtensionality
A, B:Type
f:A -> B
g1, g2:B -> A
Hg1f:forall a : A, g1 (f a) = a
Hfg1:forall b : B, f (g1 b) = b
Hg2f:forall a : A, g2 (f a) = a
Hfg2:forall b : B, f (g2 b) = b

forall x : B, g1 x = g2 x
intros; congruence. Qed.

UniqueInverse -> EqDeltaProjs

UniqueInverse -> EqDeltaProjs
UniqInv:UniqueInverse
A:Type

pi1 = pi2
apply UniqInv with delta; [apply diagonal_inverse1 | apply diagonal_inverse2]. Qed.

FunctionalExtensionality <-> UniqueInverse

FunctionalExtensionality <-> UniqueInverse

FunctionalExtensionality -> UniqueInverse

UniqueInverse -> FunctionalExtensionality

FunctionalExtensionality -> UniqueInverse
apply FunctExt_UniqInverse.

UniqueInverse -> FunctionalExtensionality
intro; apply FunctExt_iff_EqDeltaProjs, UniqInverse_EqDeltaProjs; trivial. Qed. (**********************************************************************)

Functional extensionality <-> Bijectivity of bijective composition


FunctionalExtensionality -> BijectivityBijectiveComp

FunctionalExtensionality -> BijectivityBijectiveComp
FunExt:FunctionalExtensionality
A, B, C:Type
f:A -> B
g:B -> A
Hgf:forall a : A, g (f a) = a
Hfg:forall b : B, f (g b) = b

is_inverse (action f) (action g)
FunExt:FunctionalExtensionality
A, B, C:Type
f:A -> B
g:B -> A
Hgf:forall a : A, g (f a) = a
Hfg:forall b : B, f (g b) = b

forall a : B -> C, (fun x : B => a (f (g x))) = a
FunExt:FunctionalExtensionality
A, B, C:Type
f:A -> B
g:B -> A
Hgf:forall a : A, g (f a) = a
Hfg:forall b : B, f (g b) = b
forall b : A -> C, (fun x : A => b (g (f x))) = b
FunExt:FunctionalExtensionality
A, B, C:Type
f:A -> B
g:B -> A
Hgf:forall a : A, g (f a) = a
Hfg:forall b : B, f (g b) = b

forall a : B -> C, (fun x : B => a (f (g x))) = a
intros h; apply FunExt; intro b; rewrite Hfg; reflexivity.
FunExt:FunctionalExtensionality
A, B, C:Type
f:A -> B
g:B -> A
Hgf:forall a : A, g (f a) = a
Hfg:forall b : B, f (g b) = b

forall b : A -> C, (fun x : A => b (g (f x))) = b
intros h; apply FunExt; intro a; rewrite Hgf; reflexivity. Qed.

BijectivityBijectiveComp -> FunctionalExtensionality

BijectivityBijectiveComp -> FunctionalExtensionality
BijComp:BijectivityBijectiveComp

FunctionalExtensionality
BijComp:BijectivityBijectiveComp

UniqueInverse
BijComp:BijectivityBijectiveComp
A, B:Type
f:A -> B
g1, g2:B -> A
H1:is_inverse f g1
H2:is_inverse f g2

g1 = g2
BijComp:BijectivityBijectiveComp
A, B:Type
f:A -> B
g1, g2:B -> A
H1:is_inverse f g1
H2:is_inverse f g2
Hg2f:forall a : B -> A, action g2 (action f a) = a

g1 = g2
BijComp:BijectivityBijectiveComp
A, B:Type
f:A -> B
g1, g2:B -> A
H1:is_inverse f g1
H2:is_inverse f g2
Hg2f:forall a : B -> A, action g2 (action f a) = a
Hfg1:forall b : A -> A, action f (action g1 b) = b

g1 = g2
BijComp:BijectivityBijectiveComp
A, B:Type
f:A -> B
g1, g2:B -> A
H1:is_inverse f g1
H2:is_inverse f g2
Hg2f:forall a : B -> A, action g2 (action f a) = a
Hfg1:forall b : A -> A, action f (action g1 b) = b

action g2 (action f g1) = g2
BijComp:BijectivityBijectiveComp
A, B:Type
f:A -> B
g1, g2:B -> A
H1:is_inverse f g1
H2:is_inverse f g2
Hg2f:forall a : B -> A, action g2 (action f a) = a
Hfg1:forall b : A -> A, action f (action g1 b) = b

action g2 (action f (action g1 (fun x : A => x))) = g2
BijComp:BijectivityBijectiveComp
A, B:Type
f:A -> B
g1, g2:B -> A
H1:is_inverse f g1
H2:is_inverse f g2
Hg2f:forall a : B -> A, action g2 (action f a) = a
Hfg1:forall b : A -> A, action f (action g1 b) = b

action g2 (fun x : A => x) = g2
reflexivity. Qed.