Library Coq.Logic.ExtensionalityFacts
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
- Functional extensionality
- Equality of projections from diagonal
- Unicity of inverse bijections
- Bijectivity of bijective composition
Set Implicit Arguments.
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} _.
Lemma diagonal_projs_same_behavior : forall A (x:Delta A), pi1 x = pi2 x.
Lemma diagonal_inverse1 : forall A, is_inverse (A:=A) delta pi1.
Lemma diagonal_inverse2 : forall A, is_inverse (A:=A) delta pi2.
Functional extensionality
Local Notation FunctionalExtensionality :=
(forall A B (f g : A -> B), (forall x, f x = g x) -> f = g).
Equality of projections from diagonal
Unicity of bijection inverse
Local 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)).
Local Notation BijectivityBijectiveComp := (forall A B C (f:A->B) g,
is_inverse f g -> is_inverse (A:=B->C) (action f) (action g)).