2014-08-01 01:40:09 +00:00
|
|
|
----------------------------------------------------------------------------------------------------
|
2014-07-12 08:44:46 +00:00
|
|
|
-- Copyright (c) 2014 Microsoft Corporation. All rights reserved.
|
|
|
|
-- Released under Apache 2.0 license as described in the file LICENSE.
|
|
|
|
-- Author: Leonardo de Moura
|
2014-08-01 01:40:09 +00:00
|
|
|
----------------------------------------------------------------------------------------------------
|
|
|
|
|
|
|
|
import logic.classes.inhabited logic.connectives.cast
|
2014-07-12 08:44:46 +00:00
|
|
|
|
2014-08-20 02:32:44 +00:00
|
|
|
using inhabited
|
|
|
|
|
2014-07-12 08:44:46 +00:00
|
|
|
-- Pi extensionality
|
2014-08-01 01:40:09 +00:00
|
|
|
axiom piext {A : Type} {B B' : A → Type} {H : inhabited (Π x, B x)} :
|
|
|
|
(Π x, B x) = (Π x, B' x) → B = B'
|
2014-07-12 08:44:46 +00:00
|
|
|
|
2014-08-01 01:40:09 +00:00
|
|
|
theorem cast_app {A : Type} {B B' : A → Type} (H : (Π x, B x) = (Π x, B' x)) (f : Π x, B x)
|
|
|
|
(a : A) : cast H f a == f a :=
|
2014-08-20 02:32:44 +00:00
|
|
|
have Hi [fact] : inhabited (Π x, B x), from inhabited_mk f,
|
2014-07-29 02:58:57 +00:00
|
|
|
have Hb : B = B', from piext H,
|
|
|
|
cast_app' Hb f a
|
2014-07-12 08:44:46 +00:00
|
|
|
|
2014-08-20 02:32:44 +00:00
|
|
|
theorem hcongr_fun {A : Type} {B B' : A → Type} {f : Π x, B x} {f' : Π x, B' x} (a : A)
|
2014-08-01 01:40:09 +00:00
|
|
|
(H : f == f') : f a == f' a :=
|
2014-08-20 02:32:44 +00:00
|
|
|
have Hi [fact] : inhabited (Π x, B x), from inhabited_mk f,
|
2014-07-29 02:58:57 +00:00
|
|
|
have Hb : B = B', from piext (type_eq H),
|
2014-08-20 02:32:44 +00:00
|
|
|
hcongr_fun' a H Hb
|
2014-07-12 08:44:46 +00:00
|
|
|
|
2014-08-01 01:40:09 +00:00
|
|
|
theorem hcongr {A A' : Type} {B : A → Type} {B' : A' → Type}
|
|
|
|
{f : Π x, B x} {f' : Π x, B' x} {a : A} {a' : A'}
|
|
|
|
(Hff' : f == f') (Haa' : a == a') : f a == f' a' :=
|
2014-07-29 02:58:57 +00:00
|
|
|
have H1 : ∀ (B B' : A → Type) (f : Π x, B x) (f' : Π x, B' x), f == f' → f a == f' a, from
|
2014-08-20 02:32:44 +00:00
|
|
|
take B B' f f' e, hcongr_fun a e,
|
2014-08-01 01:40:09 +00:00
|
|
|
have H2 : ∀ (B : A → Type) (B' : A' → Type) (f : Π x, B x) (f' : Π x, B' x),
|
|
|
|
f == f' → f a == f' a', from hsubst Haa' H1,
|
2014-07-29 02:58:57 +00:00
|
|
|
H2 B B' f f' Hff'
|