2014-12-12 04:14:53 +00:00
|
|
|
|
-- Copyright (c) 2014 Microsoft Corporation. All rights reserved.
|
|
|
|
|
-- Released under Apache 2.0 license as described in the file LICENSE.
|
|
|
|
|
-- Author: Jeremy Avigad, Jakob von Raumer
|
|
|
|
|
-- Ported from Coq HoTT
|
2014-12-12 18:17:50 +00:00
|
|
|
|
prelude
|
|
|
|
|
import ..path ..equiv
|
|
|
|
|
open eq
|
2014-12-12 04:14:53 +00:00
|
|
|
|
|
|
|
|
|
-- Funext
|
|
|
|
|
-- ------
|
|
|
|
|
|
|
|
|
|
-- Define function extensionality as a type class
|
2015-02-21 00:30:32 +00:00
|
|
|
|
structure funext [class] : Type :=
|
|
|
|
|
(elim : Π (A : Type) (P : A → Type ) (f g : Π x, P x), is_equiv (@apD10 A P f g))
|
2014-12-12 04:14:53 +00:00
|
|
|
|
|
|
|
|
|
|
2015-02-21 00:30:32 +00:00
|
|
|
|
namespace funext
|
2014-12-12 04:14:53 +00:00
|
|
|
|
|
2015-02-21 00:30:32 +00:00
|
|
|
|
attribute elim [instance]
|
2014-12-12 04:14:53 +00:00
|
|
|
|
|
2015-02-21 00:30:32 +00:00
|
|
|
|
definition eq_of_homotopy [F : funext] {A : Type} {P : A → Type} {f g : Π x, P x} : f ∼ g → f = g :=
|
2014-12-12 04:14:53 +00:00
|
|
|
|
is_equiv.inv (@apD10 A P f g)
|
|
|
|
|
|
2015-02-21 00:30:32 +00:00
|
|
|
|
definition eq_of_homotopy2 [F : funext] {A B : Type} {P : A → B → Type}
|
2014-12-12 18:17:50 +00:00
|
|
|
|
(f g : Πx y, P x y) : (Πx y, f x y = g x y) → f = g :=
|
2015-02-21 00:30:32 +00:00
|
|
|
|
λ E, eq_of_homotopy (λx, eq_of_homotopy (E x))
|
2014-12-12 04:14:53 +00:00
|
|
|
|
|
|
|
|
|
end funext
|