2015-04-04 04:20:19 +00:00
|
|
|
/-
|
|
|
|
Copyright (c) 2015 Floris van Doorn. All rights reserved.
|
|
|
|
Released under Apache 2.0 license as described in the file LICENSE.
|
|
|
|
|
|
|
|
Module: hit.pushout
|
|
|
|
Authors: Floris van Doorn
|
|
|
|
|
2015-04-10 01:45:18 +00:00
|
|
|
Declaration of the pushout
|
2015-04-04 04:20:19 +00:00
|
|
|
-/
|
|
|
|
|
2015-04-11 00:33:33 +00:00
|
|
|
import .type_quotient
|
2015-04-07 01:01:08 +00:00
|
|
|
|
2015-04-19 21:56:24 +00:00
|
|
|
open type_quotient eq sum equiv
|
2015-04-04 04:20:19 +00:00
|
|
|
|
|
|
|
namespace pushout
|
2015-04-23 22:27:56 +00:00
|
|
|
section
|
2015-04-04 04:20:19 +00:00
|
|
|
|
2015-04-11 00:33:33 +00:00
|
|
|
parameters {TL BL TR : Type} (f : TL → BL) (g : TL → TR)
|
|
|
|
|
|
|
|
local abbreviation A := BL + TR
|
|
|
|
inductive pushout_rel : A → A → Type :=
|
|
|
|
| Rmk : Π(x : TL), pushout_rel (inl (f x)) (inr (g x))
|
|
|
|
open pushout_rel
|
|
|
|
local abbreviation R := pushout_rel
|
|
|
|
|
|
|
|
definition pushout : Type := type_quotient pushout_rel -- TODO: define this in root namespace
|
2015-04-04 04:20:19 +00:00
|
|
|
|
2015-04-10 01:45:18 +00:00
|
|
|
definition inl (x : BL) : pushout :=
|
2015-04-11 00:33:33 +00:00
|
|
|
class_of R (inl x)
|
2015-04-04 04:20:19 +00:00
|
|
|
|
|
|
|
definition inr (x : TR) : pushout :=
|
2015-04-11 00:33:33 +00:00
|
|
|
class_of R (inr x)
|
2015-04-04 04:20:19 +00:00
|
|
|
|
|
|
|
definition glue (x : TL) : inl (f x) = inr (g x) :=
|
2015-04-11 00:33:33 +00:00
|
|
|
eq_of_rel (Rmk f g x)
|
2015-04-04 04:20:19 +00:00
|
|
|
|
|
|
|
protected definition rec {P : pushout → Type} (Pinl : Π(x : BL), P (inl x))
|
|
|
|
(Pinr : Π(x : TR), P (inr x)) (Pglue : Π(x : TL), glue x ▹ Pinl (f x) = Pinr (g x))
|
|
|
|
(y : pushout) : P y :=
|
|
|
|
begin
|
2015-04-11 00:33:33 +00:00
|
|
|
fapply (type_quotient.rec_on y),
|
|
|
|
{ intro a, cases a,
|
2015-04-04 04:20:19 +00:00
|
|
|
apply Pinl,
|
|
|
|
apply Pinr},
|
2015-04-11 00:33:33 +00:00
|
|
|
{ intros [a, a', H], cases H, apply Pglue}
|
2015-04-04 04:20:19 +00:00
|
|
|
end
|
|
|
|
|
2015-04-07 01:01:08 +00:00
|
|
|
protected definition rec_on [reducible] {P : pushout → Type} (y : pushout)
|
|
|
|
(Pinl : Π(x : BL), P (inl x)) (Pinr : Π(x : TR), P (inr x))
|
|
|
|
(Pglue : Π(x : TL), glue x ▹ Pinl (f x) = Pinr (g x)) : P y :=
|
2015-04-04 04:20:19 +00:00
|
|
|
rec Pinl Pinr Pglue y
|
|
|
|
|
2015-04-10 01:45:18 +00:00
|
|
|
--these definitions are needed until we have them definitionally
|
2015-04-07 01:01:08 +00:00
|
|
|
definition rec_inl {P : pushout → Type} (Pinl : Π(x : BL), P (inl x))
|
2015-04-04 04:20:19 +00:00
|
|
|
(Pinr : Π(x : TR), P (inr x)) (Pglue : Π(x : TL), glue x ▹ Pinl (f x) = Pinr (g x))
|
|
|
|
(x : BL) : rec Pinl Pinr Pglue (inl x) = Pinl x :=
|
2015-04-23 22:27:56 +00:00
|
|
|
idp
|
2015-04-04 04:20:19 +00:00
|
|
|
|
2015-04-07 01:01:08 +00:00
|
|
|
definition rec_inr {P : pushout → Type} (Pinl : Π(x : BL), P (inl x))
|
2015-04-04 04:20:19 +00:00
|
|
|
(Pinr : Π(x : TR), P (inr x)) (Pglue : Π(x : TL), glue x ▹ Pinl (f x) = Pinr (g x))
|
|
|
|
(x : TR) : rec Pinl Pinr Pglue (inr x) = Pinr x :=
|
2015-04-23 22:27:56 +00:00
|
|
|
idp
|
2015-04-07 01:01:08 +00:00
|
|
|
|
2015-04-19 21:56:24 +00:00
|
|
|
definition rec_glue {P : pushout → Type} (Pinl : Π(x : BL), P (inl x))
|
|
|
|
(Pinr : Π(x : TR), P (inr x)) (Pglue : Π(x : TL), glue x ▹ Pinl (f x) = Pinr (g x))
|
2015-04-23 22:27:56 +00:00
|
|
|
(x : TL) : apD (rec Pinl Pinr Pglue) (glue x) = Pglue x :=
|
2015-04-19 21:56:24 +00:00
|
|
|
sorry
|
|
|
|
|
2015-04-07 01:01:08 +00:00
|
|
|
protected definition elim {P : Type} (Pinl : BL → P) (Pinr : TR → P)
|
|
|
|
(Pglue : Π(x : TL), Pinl (f x) = Pinr (g x)) (y : pushout) : P :=
|
|
|
|
rec Pinl Pinr (λx, !tr_constant ⬝ Pglue x) y
|
|
|
|
|
2015-04-19 21:56:24 +00:00
|
|
|
protected definition elim_on [reducible] {P : Type} (y : pushout) (Pinl : BL → P)
|
2015-04-07 01:01:08 +00:00
|
|
|
(Pinr : TR → P) (Pglue : Π(x : TL), Pinl (f x) = Pinr (g x)) : P :=
|
|
|
|
elim Pinl Pinr Pglue y
|
2015-04-04 04:20:19 +00:00
|
|
|
|
2015-04-07 01:01:08 +00:00
|
|
|
definition elim_glue {P : Type} (Pinl : BL → P) (Pinr : TR → P)
|
|
|
|
(Pglue : Π(x : TL), Pinl (f x) = Pinr (g x)) (y : pushout) (x : TL)
|
2015-04-23 22:27:56 +00:00
|
|
|
: ap (elim Pinl Pinr Pglue) (glue x) = Pglue x :=
|
2015-04-07 01:01:08 +00:00
|
|
|
sorry
|
2015-04-04 04:20:19 +00:00
|
|
|
|
2015-04-19 21:56:24 +00:00
|
|
|
protected definition elim_type (Pinl : BL → Type) (Pinr : TR → Type)
|
|
|
|
(Pglue : Π(x : TL), Pinl (f x) ≃ Pinr (g x)) (y : pushout) : Type :=
|
|
|
|
elim Pinl Pinr (λx, ua (Pglue x)) y
|
|
|
|
|
|
|
|
protected definition elim_type_on [reducible] (y : pushout) (Pinl : BL → Type)
|
|
|
|
(Pinr : TR → Type) (Pglue : Π(x : TL), Pinl (f x) ≃ Pinr (g x)) : Type :=
|
|
|
|
elim_type Pinl Pinr Pglue y
|
|
|
|
|
|
|
|
definition elim_type_glue (Pinl : BL → Type) (Pinr : TR → Type)
|
|
|
|
(Pglue : Π(x : TL), Pinl (f x) ≃ Pinr (g x)) (y : pushout) (x : TL)
|
|
|
|
: transport (elim_type Pinl Pinr Pglue) (glue x) = sorry /-Pglue x-/ :=
|
|
|
|
sorry
|
2015-04-04 04:20:19 +00:00
|
|
|
|
2015-04-07 01:01:08 +00:00
|
|
|
end
|
2015-04-04 04:20:19 +00:00
|
|
|
|
2015-04-11 00:33:33 +00:00
|
|
|
open pushout equiv is_equiv unit bool
|
2015-04-07 01:01:08 +00:00
|
|
|
|
|
|
|
namespace test
|
|
|
|
definition unit_of_empty (u : empty) : unit := star
|
|
|
|
|
|
|
|
example : pushout unit_of_empty unit_of_empty ≃ bool :=
|
|
|
|
begin
|
|
|
|
fapply equiv.MK,
|
|
|
|
{ intro x, fapply (pushout.rec_on _ _ x),
|
|
|
|
intro u, exact ff,
|
|
|
|
intro u, exact tt,
|
|
|
|
intro c, cases c},
|
|
|
|
{ intro b, cases b,
|
|
|
|
exact (inl _ _ ⋆),
|
|
|
|
exact (inr _ _ ⋆)},
|
2015-04-23 22:27:56 +00:00
|
|
|
{ intro b, cases b, esimp, esimp},
|
2015-04-07 01:01:08 +00:00
|
|
|
{ intro x, fapply (pushout.rec_on _ _ x),
|
|
|
|
intro u, cases u, rewrite [↑function.compose,↑pushout.rec_on,rec_inl],
|
|
|
|
intro u, cases u, rewrite [↑function.compose,↑pushout.rec_on,rec_inr],
|
|
|
|
intro c, cases c},
|
|
|
|
end
|
|
|
|
|
|
|
|
end test
|
|
|
|
end pushout
|