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
- Functional extensionality
- Equality of projections from diagonal
- Unicity of inverse bijections
- Bijectivity of bijective composition
Set Implicit Arguments. (**********************************************************************)
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 xdestruct x as (a1,a2,Heq); assumption. Qed.forall (A : Type) (x : Delta A), pi1 x = pi2 xforall A : Type, is_inverse delta pi1split; [trivial|]; destruct b as (a1,a2,[]); reflexivity. Qed.forall A : Type, is_inverse delta pi1forall A : Type, is_inverse delta pi2split; [trivial|]; destruct b as (a1,a2,[]); reflexivity. Qed.forall A : Type, is_inverse delta pi2
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)). (**********************************************************************)
FunctionalExtensionality <-> EqDeltaProjsFunctionalExtensionality <-> EqDeltaProjsFunctionalExtensionality -> EqDeltaProjsEqDeltaProjs -> FunctionalExtensionalityintros FunExt *; apply FunExt, diagonal_projs_same_behavior.FunctionalExtensionality -> EqDeltaProjsEqDeltaProjs -> FunctionalExtensionalityrewrite EqProjs; reflexivity. Qed. (**********************************************************************)EqProjs:EqDeltaProjsA, B:Typef, g:A -> BH:forall x : A, f x = g x(fun x : A => pi1 {| pi1 := f x; pi2 := g x; eq := H x |}) = g
FunctionalExtensionality -> UniqueInverseFunctionalExtensionality -> UniqueInverseFunExt:FunctionalExtensionalityA, B:Typef:A -> Bg1, g2:B -> AHg1f:forall a : A, g1 (f a) = aHfg1:forall b : B, f (g1 b) = bHg2f:forall a : A, g2 (f a) = aHfg2:forall b : B, f (g2 b) = bg1 = g2intros; congruence. Qed.FunExt:FunctionalExtensionalityA, B:Typef:A -> Bg1, g2:B -> AHg1f:forall a : A, g1 (f a) = aHfg1:forall b : B, f (g1 b) = bHg2f:forall a : A, g2 (f a) = aHfg2:forall b : B, f (g2 b) = bforall x : B, g1 x = g2 xUniqueInverse -> EqDeltaProjsUniqueInverse -> EqDeltaProjsapply UniqInv with delta; [apply diagonal_inverse1 | apply diagonal_inverse2]. Qed.UniqInv:UniqueInverseA:Typepi1 = pi2FunctionalExtensionality <-> UniqueInverseFunctionalExtensionality <-> UniqueInverseFunctionalExtensionality -> UniqueInverseUniqueInverse -> FunctionalExtensionalityapply FunctExt_UniqInverse.FunctionalExtensionality -> UniqueInverseintro; apply FunctExt_iff_EqDeltaProjs, UniqInverse_EqDeltaProjs; trivial. Qed. (**********************************************************************)UniqueInverse -> FunctionalExtensionality
FunctionalExtensionality -> BijectivityBijectiveCompFunctionalExtensionality -> BijectivityBijectiveCompFunExt:FunctionalExtensionalityA, B, C:Typef:A -> Bg:B -> AHgf:forall a : A, g (f a) = aHfg:forall b : B, f (g b) = bis_inverse (action f) (action g)FunExt:FunctionalExtensionalityA, B, C:Typef:A -> Bg:B -> AHgf:forall a : A, g (f a) = aHfg:forall b : B, f (g b) = bforall a : B -> C, (fun x : B => a (f (g x))) = aFunExt:FunctionalExtensionalityA, B, C:Typef:A -> Bg:B -> AHgf:forall a : A, g (f a) = aHfg:forall b : B, f (g b) = bforall b : A -> C, (fun x : A => b (g (f x))) = bintros h; apply FunExt; intro b; rewrite Hfg; reflexivity.FunExt:FunctionalExtensionalityA, B, C:Typef:A -> Bg:B -> AHgf:forall a : A, g (f a) = aHfg:forall b : B, f (g b) = bforall a : B -> C, (fun x : B => a (f (g x))) = aintros h; apply FunExt; intro a; rewrite Hgf; reflexivity. Qed.FunExt:FunctionalExtensionalityA, B, C:Typef:A -> Bg:B -> AHgf:forall a : A, g (f a) = aHfg:forall b : B, f (g b) = bforall b : A -> C, (fun x : A => b (g (f x))) = bBijectivityBijectiveComp -> FunctionalExtensionalityBijectivityBijectiveComp -> FunctionalExtensionalityBijComp:BijectivityBijectiveCompFunctionalExtensionalityBijComp:BijectivityBijectiveCompUniqueInverseBijComp:BijectivityBijectiveCompA, B:Typef:A -> Bg1, g2:B -> AH1:is_inverse f g1H2:is_inverse f g2g1 = g2BijComp:BijectivityBijectiveCompA, B:Typef:A -> Bg1, g2:B -> AH1:is_inverse f g1H2:is_inverse f g2Hg2f:forall a : B -> A, action g2 (action f a) = ag1 = g2BijComp:BijectivityBijectiveCompA, B:Typef:A -> Bg1, g2:B -> AH1:is_inverse f g1H2:is_inverse f g2Hg2f:forall a : B -> A, action g2 (action f a) = aHfg1:forall b : A -> A, action f (action g1 b) = bg1 = g2BijComp:BijectivityBijectiveCompA, B:Typef:A -> Bg1, g2:B -> AH1:is_inverse f g1H2:is_inverse f g2Hg2f:forall a : B -> A, action g2 (action f a) = aHfg1:forall b : A -> A, action f (action g1 b) = baction g2 (action f g1) = g2BijComp:BijectivityBijectiveCompA, B:Typef:A -> Bg1, g2:B -> AH1:is_inverse f g1H2:is_inverse f g2Hg2f:forall a : B -> A, action g2 (action f a) = aHfg1:forall b : A -> A, action f (action g1 b) = baction g2 (action f (action g1 (fun x : A => x))) = g2reflexivity. Qed.BijComp:BijectivityBijectiveCompA, B:Typef:A -> Bg1, g2:B -> AH1:is_inverse f g1H2:is_inverse f g2Hg2f:forall a : B -> A, action g2 (action f a) = aHfg1:forall b : A -> A, action f (action g1 b) = baction g2 (fun x : A => x) = g2