2015-06-23 18:46:55 +00:00
|
|
|
/-
|
|
|
|
Copyright (c) 2015 Floris van Doorn. All rights reserved.
|
|
|
|
Released under Apache 2.0 license as described in the file LICENSE.
|
|
|
|
Author: Floris van Doorn
|
|
|
|
|
|
|
|
The "equivalence closure" of a type-valued relation.
|
2015-09-28 04:38:35 +00:00
|
|
|
A more appropriate intuition is the type of words formed from the relation,
|
|
|
|
and inverses, concatenations and reflexivity
|
2015-06-23 18:46:55 +00:00
|
|
|
-/
|
|
|
|
|
2015-11-22 06:37:13 +00:00
|
|
|
import algebra.relation eq2 arity cubical.pathover2
|
2015-06-23 18:46:55 +00:00
|
|
|
|
2016-04-11 17:11:59 +00:00
|
|
|
open eq equiv function
|
2015-06-23 18:46:55 +00:00
|
|
|
|
|
|
|
inductive e_closure {A : Type} (R : A → A → Type) : A → A → Type :=
|
|
|
|
| of_rel : Π{a a'} (r : R a a'), e_closure R a a'
|
2015-11-16 21:23:06 +00:00
|
|
|
| of_path : Π{a a'} (pp : a = a'), e_closure R a a'
|
2015-06-23 18:46:55 +00:00
|
|
|
| symm : Π{a a'} (r : e_closure R a a'), e_closure R a' a
|
|
|
|
| trans : Π{a a' a''} (r : e_closure R a a') (r' : e_closure R a' a''), e_closure R a a''
|
|
|
|
|
|
|
|
namespace e_closure
|
2015-10-01 19:52:28 +00:00
|
|
|
infix ` ⬝r `:75 := e_closure.trans
|
2015-06-23 18:46:55 +00:00
|
|
|
postfix `⁻¹ʳ`:(max+10) := e_closure.symm
|
|
|
|
notation `[`:max a `]`:0 := e_closure.of_rel a
|
2015-11-22 08:08:31 +00:00
|
|
|
notation `<`:max p `>`:0 := e_closure.of_path _ p
|
2016-03-08 05:16:45 +00:00
|
|
|
abbreviation rfl [constructor] {A : Type} {R : A → A → Type} {a : A} := of_path R (idpath a)
|
2015-06-23 18:46:55 +00:00
|
|
|
end e_closure
|
2015-11-22 06:37:13 +00:00
|
|
|
open e_closure
|
2015-06-23 18:46:55 +00:00
|
|
|
namespace relation
|
|
|
|
|
|
|
|
section
|
|
|
|
parameters {A : Type}
|
2016-03-08 05:16:45 +00:00
|
|
|
{R : A → A → Type}
|
2015-06-23 18:46:55 +00:00
|
|
|
local abbreviation T := e_closure R
|
|
|
|
|
2015-09-10 22:32:52 +00:00
|
|
|
variables ⦃a a' a'' : A⦄ {s : R a a'} {r : T a a} {B C : Type}
|
2016-03-08 05:16:45 +00:00
|
|
|
|
2015-09-10 22:32:52 +00:00
|
|
|
protected definition e_closure.elim [unfold 8] {f : A → B}
|
2015-06-23 18:46:55 +00:00
|
|
|
(e : Π⦃a a' : A⦄, R a a' → f a = f a') (t : T a a') : f a = f a' :=
|
|
|
|
begin
|
2015-11-16 21:23:06 +00:00
|
|
|
induction t with a a' r a a' pp a a' r IH a a' a'' r r' IH₁ IH₂,
|
2015-06-23 18:46:55 +00:00
|
|
|
exact e r,
|
2015-11-16 21:23:06 +00:00
|
|
|
exact ap f pp,
|
|
|
|
exact IH⁻¹,
|
|
|
|
exact IH₁ ⬝ IH₂
|
2015-06-23 18:46:55 +00:00
|
|
|
end
|
|
|
|
|
2015-07-29 12:17:16 +00:00
|
|
|
definition ap_e_closure_elim_h [unfold 12] {B C : Type} {f : A → B} {g : B → C}
|
2015-06-23 18:46:55 +00:00
|
|
|
(e : Π⦃a a' : A⦄, R a a' → f a = f a')
|
|
|
|
{e' : Π⦃a a' : A⦄, R a a' → g (f a) = g (f a')}
|
|
|
|
(p : Π⦃a a' : A⦄ (s : R a a'), ap g (e s) = e' s) (t : T a a')
|
|
|
|
: ap g (e_closure.elim e t) = e_closure.elim e' t :=
|
|
|
|
begin
|
2015-11-16 21:23:06 +00:00
|
|
|
induction t with a a' r a a' pp a a' r IH a a' a'' r r' IH₁ IH₂,
|
2015-06-23 18:46:55 +00:00
|
|
|
apply p,
|
2015-11-16 21:23:06 +00:00
|
|
|
induction pp, reflexivity,
|
|
|
|
exact ap_inv g (e_closure.elim e r) ⬝ inverse2 IH,
|
|
|
|
exact ap_con g (e_closure.elim e r) (e_closure.elim e r') ⬝ (IH₁ ◾ IH₂)
|
2015-06-23 18:46:55 +00:00
|
|
|
end
|
|
|
|
|
2016-04-11 17:11:59 +00:00
|
|
|
definition ap_e_closure_elim_h_symm [unfold_full] {B C : Type} {f : A → B} {g : B → C}
|
|
|
|
{e : Π⦃a a' : A⦄, R a a' → f a = f a'}
|
|
|
|
{e' : Π⦃a a' : A⦄, R a a' → g (f a) = g (f a')}
|
|
|
|
(p : Π⦃a a' : A⦄ (s : R a a'), ap g (e s) = e' s) (t : T a a') :
|
|
|
|
ap_e_closure_elim_h e p t⁻¹ʳ = ap_inv g (e_closure.elim e t) ⬝ (ap_e_closure_elim_h e p t)⁻² :=
|
|
|
|
by reflexivity
|
|
|
|
|
|
|
|
definition ap_e_closure_elim_h_trans [unfold_full] {B C : Type} {f : A → B} {g : B → C}
|
|
|
|
{e : Π⦃a a' : A⦄, R a a' → f a = f a'}
|
|
|
|
{e' : Π⦃a a' : A⦄, R a a' → g (f a) = g (f a')}
|
|
|
|
(p : Π⦃a a' : A⦄ (s : R a a'), ap g (e s) = e' s) (t : T a a') (t' : T a' a'')
|
|
|
|
: ap_e_closure_elim_h e p (t ⬝r t') = ap_con g (e_closure.elim e t) (e_closure.elim e t') ⬝
|
|
|
|
(ap_e_closure_elim_h e p t ◾ ap_e_closure_elim_h e p t') :=
|
|
|
|
by reflexivity
|
|
|
|
|
2015-11-16 21:23:06 +00:00
|
|
|
definition ap_e_closure_elim [unfold 10] {B C : Type} {f : A → B} (g : B → C)
|
2015-11-22 06:37:13 +00:00
|
|
|
(e : Π⦃a a' : A⦄, R a a' → f a = f a') (t : T a a')
|
2015-06-23 18:46:55 +00:00
|
|
|
: ap g (e_closure.elim e t) = e_closure.elim (λa a' r, ap g (e r)) t :=
|
|
|
|
ap_e_closure_elim_h e (λa a' s, idp) t
|
|
|
|
|
2016-04-11 17:11:59 +00:00
|
|
|
definition ap_e_closure_elim_symm [unfold_full] {B C : Type} {f : A → B} (g : B → C)
|
2015-11-22 06:37:13 +00:00
|
|
|
(e : Π⦃a a' : A⦄, R a a' → f a = f a') (t : T a a')
|
|
|
|
: ap_e_closure_elim g e t⁻¹ʳ = ap_inv g (e_closure.elim e t) ⬝ (ap_e_closure_elim g e t)⁻² :=
|
|
|
|
by reflexivity
|
|
|
|
|
2016-04-11 17:11:59 +00:00
|
|
|
definition ap_e_closure_elim_trans [unfold_full] {B C : Type} {f : A → B} (g : B → C)
|
2015-11-22 06:37:13 +00:00
|
|
|
(e : Π⦃a a' : A⦄, R a a' → f a = f a') (t : T a a') (t' : T a' a'')
|
|
|
|
: ap_e_closure_elim g e (t ⬝r t') = ap_con g (e_closure.elim e t) (e_closure.elim e t') ⬝
|
|
|
|
(ap_e_closure_elim g e t ◾ ap_e_closure_elim g e t') :=
|
|
|
|
by reflexivity
|
|
|
|
|
2016-04-11 17:11:59 +00:00
|
|
|
definition e_closure_elim_eq [unfold 8] {f : A → B}
|
|
|
|
{e e' : Π⦃a a' : A⦄, R a a' → f a = f a'} (p : e ~3 e') (t : T a a')
|
|
|
|
: e_closure.elim e t = e_closure.elim e' t :=
|
|
|
|
begin
|
|
|
|
induction t with a a' r a a' pp a a' r IH a a' a'' r r' IH₁ IH₂,
|
|
|
|
apply p,
|
|
|
|
reflexivity,
|
|
|
|
exact IH⁻²,
|
|
|
|
exact IH₁ ◾ IH₂
|
|
|
|
end
|
|
|
|
|
|
|
|
-- TODO: formulate and prove this without using function extensionality,
|
|
|
|
-- and modify the proofs using this to also not use function extensionality
|
|
|
|
-- strategy: use `e_closure_elim_eq` instead of `ap ... (eq_of_homotopy3 p)`
|
2015-06-23 18:46:55 +00:00
|
|
|
definition ap_e_closure_elim_h_eq {B C : Type} {f : A → B} {g : B → C}
|
|
|
|
(e : Π⦃a a' : A⦄, R a a' → f a = f a')
|
|
|
|
{e' : Π⦃a a' : A⦄, R a a' → g (f a) = g (f a')}
|
|
|
|
(p : Π⦃a a' : A⦄ (s : R a a'), ap g (e s) = e' s) (t : T a a')
|
|
|
|
: ap_e_closure_elim_h e p t =
|
|
|
|
ap_e_closure_elim g e t ⬝ ap (λx, e_closure.elim x t) (eq_of_homotopy3 p) :=
|
|
|
|
begin
|
|
|
|
fapply homotopy3.rec_on p,
|
|
|
|
intro q, esimp at q, induction q,
|
|
|
|
esimp, rewrite eq_of_homotopy3_id
|
|
|
|
end
|
|
|
|
|
|
|
|
theorem ap_ap_e_closure_elim_h {B C D : Type} {f : A → B}
|
|
|
|
{g : B → C} (h : C → D)
|
|
|
|
(e : Π⦃a a' : A⦄, R a a' → f a = f a')
|
|
|
|
{e' : Π⦃a a' : A⦄, R a a' → g (f a) = g (f a')}
|
|
|
|
(p : Π⦃a a' : A⦄ (s : R a a'), ap g (e s) = e' s) (t : T a a')
|
|
|
|
: square (ap (ap h) (ap_e_closure_elim_h e p t))
|
|
|
|
(ap_e_closure_elim_h e (λa a' s, ap_compose h g (e s)) t)
|
|
|
|
(ap_compose h g (e_closure.elim e t))⁻¹
|
|
|
|
(ap_e_closure_elim_h e' (λa a' s, (ap (ap h) (p s))⁻¹) t) :=
|
|
|
|
begin
|
2015-11-16 21:23:06 +00:00
|
|
|
induction t with a a' r a a' pp a a' r IH a a' a'' r r' IH₁ IH₂,
|
2015-07-29 12:17:16 +00:00
|
|
|
{ esimp,
|
2015-06-23 18:46:55 +00:00
|
|
|
apply square_of_eq, exact !con.right_inv ⬝ !con.left_inv⁻¹},
|
2015-11-16 21:23:06 +00:00
|
|
|
{ induction pp, apply ids},
|
2015-07-29 12:17:16 +00:00
|
|
|
{ rewrite [▸*,ap_con (ap h)],
|
2015-06-23 18:46:55 +00:00
|
|
|
refine (transpose !ap_compose_inv)⁻¹ᵛ ⬝h _,
|
|
|
|
rewrite [con_inv,inv_inv,-inv2_inv],
|
2015-11-16 21:23:06 +00:00
|
|
|
exact !ap_inv2 ⬝v square_inv2 IH},
|
2015-07-29 12:17:16 +00:00
|
|
|
{ rewrite [▸*,ap_con (ap h)],
|
2015-06-23 18:46:55 +00:00
|
|
|
refine (transpose !ap_compose_con)⁻¹ᵛ ⬝h _,
|
|
|
|
rewrite [con_inv,inv_inv,con2_inv],
|
2015-11-16 21:23:06 +00:00
|
|
|
refine !ap_con2 ⬝v square_con2 IH₁ IH₂},
|
2015-06-23 18:46:55 +00:00
|
|
|
end
|
|
|
|
|
|
|
|
theorem ap_ap_e_closure_elim {B C D : Type} {f : A → B}
|
|
|
|
(g : B → C) (h : C → D)
|
|
|
|
(e : Π⦃a a' : A⦄, R a a' → f a = f a') (t : T a a')
|
|
|
|
: square (ap (ap h) (ap_e_closure_elim g e t))
|
|
|
|
(ap_e_closure_elim_h e (λa a' s, ap_compose h g (e s)) t)
|
|
|
|
(ap_compose h g (e_closure.elim e t))⁻¹
|
2016-04-11 17:11:59 +00:00
|
|
|
(ap_e_closure_elim h (λa a' r, ap g (e r)) t) :=
|
2015-06-23 18:46:55 +00:00
|
|
|
!ap_ap_e_closure_elim_h
|
|
|
|
|
2016-04-11 17:11:59 +00:00
|
|
|
definition ap_e_closure_elim_h_zigzag {B C D : Type} {f : A → B}
|
|
|
|
{g : B → C} (h : C → D)
|
|
|
|
(e : Π⦃a a' : A⦄, R a a' → f a = f a')
|
|
|
|
{e' : Π⦃a a' : A⦄, R a a' → h (g (f a)) = h (g (f a'))}
|
|
|
|
(p : Π⦃a a' : A⦄ (s : R a a'), ap (h ∘ g) (e s) = e' s) (t : T a a')
|
|
|
|
: ap_e_closure_elim h (λa a' s, ap g (e s)) t ⬝
|
|
|
|
(ap_e_closure_elim_h e (λa a' s, ap_compose h g (e s)) t)⁻¹ ⬝
|
|
|
|
ap_e_closure_elim_h e p t =
|
|
|
|
ap_e_closure_elim_h (λa a' s, ap g (e s)) (λa a' s, (ap_compose h g (e s))⁻¹ ⬝ p s) t :=
|
|
|
|
begin
|
2016-11-24 05:13:05 +00:00
|
|
|
refine whisker_right _ (eq_of_square (ap_ap_e_closure_elim g h e t)⁻¹ʰ)⁻¹ ⬝ _,
|
2016-04-11 17:11:59 +00:00
|
|
|
refine !con.assoc ⬝ _, apply inv_con_eq_of_eq_con, apply eq_of_square,
|
|
|
|
apply transpose,
|
|
|
|
-- the rest of the proof is almost the same as the proof of ap_ap_e_closure_elim[_h].
|
|
|
|
-- Is there a connection between these theorems?
|
|
|
|
induction t with a a' r a a' pp a a' r IH a a' a'' r r' IH₁ IH₂,
|
|
|
|
{ esimp, apply square_of_eq, apply idp_con},
|
|
|
|
{ induction pp, apply ids},
|
|
|
|
{ rewrite [▸*,ap_con (ap h)],
|
|
|
|
refine (transpose !ap_compose_inv)⁻¹ᵛ ⬝h _,
|
|
|
|
rewrite [con_inv,inv_inv,-inv2_inv],
|
|
|
|
exact !ap_inv2 ⬝v square_inv2 IH},
|
|
|
|
{ rewrite [▸*,ap_con (ap h)],
|
|
|
|
refine (transpose !ap_compose_con)⁻¹ᵛ ⬝h _,
|
|
|
|
rewrite [con_inv,inv_inv,con2_inv],
|
|
|
|
refine !ap_con2 ⬝v square_con2 IH₁ IH₂},
|
|
|
|
|
|
|
|
end
|
|
|
|
|
2015-06-23 18:46:55 +00:00
|
|
|
definition is_equivalence_e_closure : is_equivalence T :=
|
|
|
|
begin
|
|
|
|
constructor,
|
|
|
|
intro a, exact rfl,
|
|
|
|
intro a a' t, exact t⁻¹ʳ,
|
|
|
|
intro a a' a'' t t', exact t ⬝r t',
|
|
|
|
end
|
|
|
|
|
2015-11-22 06:37:13 +00:00
|
|
|
/- dependent elimination -/
|
2015-09-10 22:32:52 +00:00
|
|
|
|
|
|
|
variables {P : B → Type} {Q : C → Type} {f : A → B} {g : B → C} {f' : Π(a : A), P (f a)}
|
2015-11-22 06:37:13 +00:00
|
|
|
protected definition e_closure.elimo [unfold 11] (p : Π⦃a a' : A⦄, R a a' → f a = f a')
|
|
|
|
(po : Π⦃a a' : A⦄ (s : R a a'), f' a =[p s] f' a') (t : T a a')
|
|
|
|
: f' a =[e_closure.elim p t] f' a' :=
|
2015-09-10 22:32:52 +00:00
|
|
|
begin
|
2015-11-16 21:23:06 +00:00
|
|
|
induction t with a a' r a a' pp a a' r IH a a' a'' r r' IH₁ IH₂,
|
2015-09-10 22:32:52 +00:00
|
|
|
exact po r,
|
2015-11-16 21:23:06 +00:00
|
|
|
induction pp, constructor,
|
|
|
|
exact IH⁻¹ᵒ,
|
|
|
|
exact IH₁ ⬝o IH₂
|
2015-09-10 22:32:52 +00:00
|
|
|
end
|
|
|
|
|
2016-04-11 17:11:59 +00:00
|
|
|
definition elimo_symm [unfold_full] (p : Π⦃a a' : A⦄, R a a' → f a = f a')
|
2015-11-22 06:37:13 +00:00
|
|
|
(po : Π⦃a a' : A⦄ (s : R a a'), f' a =[p s] f' a') (t : T a a')
|
|
|
|
: e_closure.elimo p po t⁻¹ʳ = (e_closure.elimo p po t)⁻¹ᵒ :=
|
|
|
|
by reflexivity
|
|
|
|
|
2016-04-11 17:11:59 +00:00
|
|
|
definition elimo_trans [unfold_full] (p : Π⦃a a' : A⦄, R a a' → f a = f a')
|
2015-11-22 06:37:13 +00:00
|
|
|
(po : Π⦃a a' : A⦄ (s : R a a'), f' a =[p s] f' a') (t : T a a') (t' : T a' a'')
|
|
|
|
: e_closure.elimo p po (t ⬝r t') = e_closure.elimo p po t ⬝o e_closure.elimo p po t' :=
|
|
|
|
by reflexivity
|
|
|
|
|
2015-09-10 22:32:52 +00:00
|
|
|
definition ap_e_closure_elimo_h [unfold 12] {g' : Πb, Q (g b)}
|
|
|
|
(p : Π⦃a a' : A⦄, R a a' → f a = f a')
|
|
|
|
(po : Π⦃a a' : A⦄ (s : R a a'), g' (f a) =[p s] g' (f a'))
|
2016-03-19 15:25:08 +00:00
|
|
|
(q : Π⦃a a' : A⦄ (s : R a a'), apd g' (p s) = po s)
|
|
|
|
(t : T a a') : apd g' (e_closure.elim p t) = e_closure.elimo p po t :=
|
2015-09-10 22:32:52 +00:00
|
|
|
begin
|
2015-11-16 21:23:06 +00:00
|
|
|
induction t with a a' r a a' pp a a' r IH a a' a'' r r' IH₁ IH₂,
|
2015-09-10 22:32:52 +00:00
|
|
|
apply q,
|
2015-11-16 21:23:06 +00:00
|
|
|
induction pp, reflexivity,
|
2015-09-10 22:32:52 +00:00
|
|
|
esimp [e_closure.elim],
|
2016-03-19 15:25:08 +00:00
|
|
|
exact apd_inv g' (e_closure.elim p r) ⬝ IH⁻²ᵒ,
|
|
|
|
exact apd_con g' (e_closure.elim p r) (e_closure.elim p r') ⬝ (IH₁ ◾o IH₂)
|
2015-09-10 22:32:52 +00:00
|
|
|
end
|
|
|
|
|
2015-11-22 06:37:13 +00:00
|
|
|
theorem e_closure_elimo_ap {g' : Π(a : A), Q (g (f a))}
|
2015-11-16 21:23:06 +00:00
|
|
|
(p : Π⦃a a' : A⦄, R a a' → f a = f a')
|
|
|
|
(po : Π⦃a a' : A⦄ (s : R a a'), g' a =[ap g (p s)] g' a')
|
2015-11-22 06:37:13 +00:00
|
|
|
(t : T a a') : e_closure.elimo p (λa a' s, pathover_of_pathover_ap Q g (po s)) t =
|
|
|
|
pathover_of_pathover_ap Q g (change_path (ap_e_closure_elim g p t)⁻¹
|
|
|
|
(e_closure.elimo (λa a' r, ap g (p r)) po t)) :=
|
2015-11-16 21:23:06 +00:00
|
|
|
begin
|
|
|
|
induction t with a a' r a a' pp a a' r IH a a' a'' r r' IH₁ IH₂,
|
2015-11-22 06:37:13 +00:00
|
|
|
{ reflexivity},
|
|
|
|
{ induction pp; reflexivity},
|
2016-04-11 17:11:59 +00:00
|
|
|
{ rewrite [+elimo_symm, ap_e_closure_elim_symm, IH, con_inv, change_path_con, ▸*, -inv2_inv,
|
2015-11-22 06:37:13 +00:00
|
|
|
change_path_invo, pathover_of_pathover_ap_invo]},
|
2016-04-11 17:11:59 +00:00
|
|
|
{ rewrite [+elimo_trans, ap_e_closure_elim_trans, IH₁, IH₂, con_inv, change_path_con, ▸*,
|
|
|
|
con2_inv, change_path_cono, pathover_of_pathover_ap_cono]},
|
2015-11-16 21:23:06 +00:00
|
|
|
end
|
2015-09-10 22:32:52 +00:00
|
|
|
|
2015-06-23 18:46:55 +00:00
|
|
|
end
|
|
|
|
end relation
|