From HoTT Require Import Basics Types.[Loading ML file number_string_notation_plugin.cmxs (using legacy method) ... done ]
(** Vertical composition of squares *)
Section concat_square_vert .
Context {X : Type }.
(** 0-paths *)
Context {a0 b0 c0 : X}.
Context {a1 b1 c1 : X}.
(** 1-paths *)
Context {a01 : a0 = a1}.
Context {b01 : b0 = b1}.
Context {c01 : c0 = c1}.
Context {ab0 : a0 = b0}.
Context {ab1 : a1 = b1}.
Context {bc0 : b0 = c0}.
Context {bc1 : b1 = c1}.
(** 2-paths *)
Context (p : ab0 @ b01 = a01 @ ab1).
Context (q : bc0 @ c01 = b01 @ bc1).
Local Definition concat_square_vert :
(ab0 @ bc0) @ c01 = a01 @ (ab1 @ bc1).X : Type a0, b0, c0, a1, b1, c1 : X a01 : a0 = a1 b01 : b0 = b1 c01 : c0 = c1 ab0 : a0 = b0 ab1 : a1 = b1 bc0 : b0 = c0 bc1 : b1 = c1 p : ab0 @ b01 = a01 @ ab1 q : bc0 @ c01 = b01 @ bc1
(ab0 @ bc0) @ c01 = a01 @ (ab1 @ bc1)
Proof .X : Type a0, b0, c0, a1, b1, c1 : X a01 : a0 = a1 b01 : b0 = b1 c01 : c0 = c1 ab0 : a0 = b0 ab1 : a1 = b1 bc0 : b0 = c0 bc1 : b1 = c1 p : ab0 @ b01 = a01 @ ab1 q : bc0 @ c01 = b01 @ bc1
(ab0 @ bc0) @ c01 = a01 @ (ab1 @ bc1)
refine (concat_pp_p _ _ _ @ _).X : Type a0, b0, c0, a1, b1, c1 : X a01 : a0 = a1 b01 : b0 = b1 c01 : c0 = c1 ab0 : a0 = b0 ab1 : a1 = b1 bc0 : b0 = c0 bc1 : b1 = c1 p : ab0 @ b01 = a01 @ ab1 q : bc0 @ c01 = b01 @ bc1
ab0 @ (bc0 @ c01) = a01 @ (ab1 @ bc1)
refine (whiskerL _ q @ _).X : Type a0, b0, c0, a1, b1, c1 : X a01 : a0 = a1 b01 : b0 = b1 c01 : c0 = c1 ab0 : a0 = b0 ab1 : a1 = b1 bc0 : b0 = c0 bc1 : b1 = c1 p : ab0 @ b01 = a01 @ ab1 q : bc0 @ c01 = b01 @ bc1
ab0 @ (b01 @ bc1) = a01 @ (ab1 @ bc1)
refine (concat_p_pp _ _ _ @ _).X : Type a0, b0, c0, a1, b1, c1 : X a01 : a0 = a1 b01 : b0 = b1 c01 : c0 = c1 ab0 : a0 = b0 ab1 : a1 = b1 bc0 : b0 = c0 bc1 : b1 = c1 p : ab0 @ b01 = a01 @ ab1 q : bc0 @ c01 = b01 @ bc1
(ab0 @ b01) @ bc1 = a01 @ (ab1 @ bc1)
refine (whiskerR p _ @ _).X : Type a0, b0, c0, a1, b1, c1 : X a01 : a0 = a1 b01 : b0 = b1 c01 : c0 = c1 ab0 : a0 = b0 ab1 : a1 = b1 bc0 : b0 = c0 bc1 : b1 = c1 p : ab0 @ b01 = a01 @ ab1 q : bc0 @ c01 = b01 @ bc1
(a01 @ ab1) @ bc1 = a01 @ (ab1 @ bc1)
apply concat_pp_p.
Defined .
End concat_square_vert .
Infix "[-]" := (concat_square_vert) (at level 60 ).
(** Horizontal composition of squares *)
Section concat_square_hor .
Context {X : Type }.
(** 0-paths *)
Context {a0 b0 c0 : X}.
Context {a1 b1 c1 : X}.
(** 1-paths *)
Context {a01 : a0 = a1}.
Context {b01 : b0 = b1}.
Context {c01 : c0 = c1}.
Context {ab0 : a0 = b0}.
Context {ab1 : a1 = b1}.
Context {bc0 : b0 = c0}.
Context {bc1 : b1 = c1}.
(** 2-paths *)
Context (p : a01 @ ab1 = ab0 @ b01).
Context (q : b01 @ bc1 = bc0 @ c01).
Local Definition concat_square_hor :
a01 @ (ab1 @ bc1) = (ab0 @ bc0) @ c01.X : Type a0, b0, c0, a1, b1, c1 : X a01 : a0 = a1 b01 : b0 = b1 c01 : c0 = c1 ab0 : a0 = b0 ab1 : a1 = b1 bc0 : b0 = c0 bc1 : b1 = c1 p : a01 @ ab1 = ab0 @ b01 q : b01 @ bc1 = bc0 @ c01
a01 @ (ab1 @ bc1) = (ab0 @ bc0) @ c01
Proof .X : Type a0, b0, c0, a1, b1, c1 : X a01 : a0 = a1 b01 : b0 = b1 c01 : c0 = c1 ab0 : a0 = b0 ab1 : a1 = b1 bc0 : b0 = c0 bc1 : b1 = c1 p : a01 @ ab1 = ab0 @ b01 q : b01 @ bc1 = bc0 @ c01
a01 @ (ab1 @ bc1) = (ab0 @ bc0) @ c01
refine (concat_p_pp _ _ _ @ _).X : Type a0, b0, c0, a1, b1, c1 : X a01 : a0 = a1 b01 : b0 = b1 c01 : c0 = c1 ab0 : a0 = b0 ab1 : a1 = b1 bc0 : b0 = c0 bc1 : b1 = c1 p : a01 @ ab1 = ab0 @ b01 q : b01 @ bc1 = bc0 @ c01
(a01 @ ab1) @ bc1 = (ab0 @ bc0) @ c01
refine (whiskerR p _ @ _).X : Type a0, b0, c0, a1, b1, c1 : X a01 : a0 = a1 b01 : b0 = b1 c01 : c0 = c1 ab0 : a0 = b0 ab1 : a1 = b1 bc0 : b0 = c0 bc1 : b1 = c1 p : a01 @ ab1 = ab0 @ b01 q : b01 @ bc1 = bc0 @ c01
(ab0 @ b01) @ bc1 = (ab0 @ bc0) @ c01
refine (concat_pp_p _ _ _ @ _).X : Type a0, b0, c0, a1, b1, c1 : X a01 : a0 = a1 b01 : b0 = b1 c01 : c0 = c1 ab0 : a0 = b0 ab1 : a1 = b1 bc0 : b0 = c0 bc1 : b1 = c1 p : a01 @ ab1 = ab0 @ b01 q : b01 @ bc1 = bc0 @ c01
ab0 @ (b01 @ bc1) = (ab0 @ bc0) @ c01
refine (whiskerL _ q @ _).X : Type a0, b0, c0, a1, b1, c1 : X a01 : a0 = a1 b01 : b0 = b1 c01 : c0 = c1 ab0 : a0 = b0 ab1 : a1 = b1 bc0 : b0 = c0 bc1 : b1 = c1 p : a01 @ ab1 = ab0 @ b01 q : b01 @ bc1 = bc0 @ c01
ab0 @ (bc0 @ c01) = (ab0 @ bc0) @ c01
apply concat_p_pp.
Defined .
End concat_square_hor .
Infix "[I]" := (concat_square_hor) (at level 60 ).
(** We will frequently use the following equivalences. *)
Definition rlucancel {X } {a b : X} {p q : a = b} :
(p = q) <~> (p @ 1 = 1 @ q).X : Type a, b : X p, q : a = b
p = q <~> p @ 1 = 1 @ q
Proof .X : Type a, b : X p, q : a = b
p = q <~> p @ 1 = 1 @ q
refine (equiv_compose' _ _).X : Type a, b : X p, q : a = b
?Goal <~> p @ 1 = 1 @ q
- X : Type a, b : X p, q : a = b
?Goal <~> p @ 1 = 1 @ q
exact (equiv_concat_r (concat_1p _)^ _).
- X : Type a, b : X p, q : a = b
p = q <~> p @ 1 = q
exact (equiv_concat_l (concat_p1 _) _).
Defined .
Definition rlucancel_inv {X } {a b : X} {p q : a = b} := (@rlucancel X a b p q)^-1 .
Definition lrucancel {X } {a b : X} {p q : a = b} :
(p = q) <~> (1 @ p = q @ 1 ).X : Type a, b : X p, q : a = b
p = q <~> 1 @ p = q @ 1
Proof .X : Type a, b : X p, q : a = b
p = q <~> 1 @ p = q @ 1
refine (equiv_compose' _ _).X : Type a, b : X p, q : a = b
?Goal <~> 1 @ p = q @ 1
- X : Type a, b : X p, q : a = b
?Goal <~> 1 @ p = q @ 1
exact (equiv_concat_r (concat_p1 _)^ _).
- X : Type a, b : X p, q : a = b
p = q <~> 1 @ p = q
exact (equiv_concat_l (concat_1p _) _).
Defined .
(** This special case of [equiv_path_ind] comes up a lot. *)
Definition equiv_path_ind_rlucancel {X } (a b : X) (p : a = b)
(P : forall (q : a = b), p @ 1 = 1 @ q -> Type )
(r : P p (rlucancel 1 ))
: forall (q : a = b) (s : p @ 1 = 1 @ q), P q s.X : Type a, b : X p : a = b P : forall q : a = b, p @ 1 = 1 @ q -> Type r : P p (rlucancel 1 )
forall (q : a = b) (s : p @ 1 = 1 @ q), P q s
Proof .X : Type a, b : X p : a = b P : forall q : a = b, p @ 1 = 1 @ q -> Type r : P p (rlucancel 1 )
forall (q : a = b) (s : p @ 1 = 1 @ q), P q s
snapply (equiv_path_ind (fun _ => rlucancel)). X : Type a, b : X p : a = b P : forall q : a = b, p @ 1 = 1 @ q -> Type r : P p (rlucancel 1 )
P p ((fun b0 : a = b => rlucancel) p 1 )
exact r.
Defined .
(** This special case of [equiv_path_ind] comes up a lot. *)
Definition equiv_path_ind_lrucancel {X } (a b : X) (p : a = b)
(P : forall (q : a = b), 1 @ p = q @ 1 -> Type )
(r : P p (lrucancel 1 ))
: forall (q : a = b) (s : 1 @ p = q @ 1 ), P q s.X : Type a, b : X p : a = b P : forall q : a = b, 1 @ p = q @ 1 -> Type r : P p (lrucancel 1 )
forall (q : a = b) (s : 1 @ p = q @ 1 ), P q s
Proof .X : Type a, b : X p : a = b P : forall q : a = b, 1 @ p = q @ 1 -> Type r : P p (lrucancel 1 )
forall (q : a = b) (s : 1 @ p = q @ 1 ), P q s
snapply (equiv_path_ind (fun _ => lrucancel)). X : Type a, b : X p : a = b P : forall q : a = b, 1 @ p = q @ 1 -> Type r : P p (lrucancel 1 )
P p ((fun b0 : a = b => lrucancel) p 1 )
exact r.
Defined .
(** Interaction of the above equivalences with square composition. *)
Definition rlucancel_sVs_1_pp {X } {a b c : X} {p : a = b} {q : b = c} {r } (theta : p @ q = r) :
(rlucancel 1 [-] rlucancel 1 ) @ whiskerL _ theta = whiskerR theta _ @ (rlucancel 1 ).X : Type a, b, c : X p : a = b q : b = c r : a = c theta : p @ q = r
(rlucancel 1 [-] rlucancel 1 ) @ whiskerL 1 theta =
whiskerR theta 1 @ rlucancel 1
Proof .X : Type a, b, c : X p : a = b q : b = c r : a = c theta : p @ q = r
(rlucancel 1 [-] rlucancel 1 ) @ whiskerL 1 theta =
whiskerR theta 1 @ rlucancel 1
by destruct theta, p, q.
Defined .
Definition lrucancel_sHs_1_pp {X } {a b c : X} {p : a = b} {q : b = c} {r } (theta : p @ q = r) :
(lrucancel 1 [I] lrucancel 1 ) @ whiskerR theta _ = whiskerL _ theta @ (lrucancel 1 ).X : Type a, b, c : X p : a = b q : b = c r : a = c theta : p @ q = r
(lrucancel 1 [I] lrucancel 1 ) @ whiskerR theta 1 =
whiskerL 1 theta @ lrucancel 1
Proof .X : Type a, b, c : X p : a = b q : b = c r : a = c theta : p @ q = r
(lrucancel 1 [I] lrucancel 1 ) @ whiskerR theta 1 =
whiskerL 1 theta @ lrucancel 1
by destruct theta, p, q.
Defined .
Definition rlucancel_sHs_1 {X } {a b : X} (p : a = b) :
(rlucancel 1 [I] rlucancel 1 ) = rlucancel (idpath p).X : Type a, b : X p : a = b
rlucancel 1 [I] rlucancel 1 = rlucancel 1
Proof .X : Type a, b : X p : a = b
rlucancel 1 [I] rlucancel 1 = rlucancel 1
by destruct p.
Defined .
Definition lrucancel_sVs_1 {X } {a b : X} (p : a = b) :
(lrucancel 1 [-] lrucancel 1 ) = lrucancel (idpath p).X : Type a, b : X p : a = b
lrucancel 1 [-] lrucancel 1 = lrucancel 1
Proof .X : Type a, b : X p : a = b
lrucancel 1 [-] lrucancel 1 = lrucancel 1
by destruct p.
Defined .
(** Naturality of composition with 1. *)
Definition ulnat {X } {a b : X} {u v : a = b} (p : u = v) :
whiskerL 1 p @ concat_1p v = concat_1p u @ p.X : Type a, b : X u, v : a = b p : u = v
whiskerL 1 p @ concat_1p v = concat_1p u @ p
Proof .X : Type a, b : X u, v : a = b p : u = v
whiskerL 1 p @ concat_1p v = concat_1p u @ p
destruct p.X : Type a, b : X u : a = b
whiskerL 1 1 @ concat_1p u = concat_1p u @ 1
exact (lrucancel 1 ).
Defined .
Definition urnat {X } {a b : X} {u v : a = b} (p : u = v) :
whiskerR p 1 @ concat_p1 v = concat_p1 u @ p.X : Type a, b : X u, v : a = b p : u = v
whiskerR p 1 @ concat_p1 v = concat_p1 u @ p
Proof .X : Type a, b : X u, v : a = b p : u = v
whiskerR p 1 @ concat_p1 v = concat_p1 u @ p
destruct p.X : Type a, b : X u : a = b
whiskerR 1 1 @ concat_p1 u = concat_p1 u @ 1
exact (lrucancel 1 ).
Defined .
(** Exchange law for whiskering on the left and on the right. *)
Definition wlrnat {X } {a b c : X} {u v : a = b} {x y : b = c} p q :
whiskerL u p @ whiskerR q y = whiskerR q x @ whiskerL v p.X : Type a, b, c : X u, v : a = b x, y : b = c p : x = y q : u = v
whiskerL u p @ whiskerR q y =
whiskerR q x @ whiskerL v p
Proof .X : Type a, b, c : X u, v : a = b x, y : b = c p : x = y q : u = v
whiskerL u p @ whiskerR q y =
whiskerR q x @ whiskerL v p
by destruct p, q.
Defined .
(** Eckmann-Hilton. This is also proved as [eckmann_hilton] in PathGroupoids.v, but we need this particular proof in order to prove the syllepsis. *)
Theorem eh {X } {a : X} (p q : idpath a = idpath a) :
p @ q = q @ p.X : Type a : X p, q : 1 = 1
p @ q = q @ p
Proof .X : Type a : X p, q : 1 = 1
p @ q = q @ p
refine (_ @ rlucancel_inv (urnat q [-] ulnat p)).X : Type a : X p, q : 1 = 1
p @ q = whiskerR q 1 @ whiskerL 1 p
refine ((rlucancel_inv (ulnat p [-] urnat q))^ @ _).X : Type a : X p, q : 1 = 1
whiskerL 1 p @ whiskerR q 1 =
whiskerR q 1 @ whiskerL 1 p
exact (wlrnat p q).
Defined .
(** Eckmann-Hilton on reflexivity. *)
Local Definition eh_1p_gen {X } {a b : X} {u v : a = b} (p : u = v) {q } (theta : whiskerR p 1 @ 1 = 1 @ q) :
(rlucancel_inv (1 [-] theta))^ @ wlrnat 1 p @ rlucancel_inv (theta [-] 1 ) @ concat_p1 q = concat_1p q.X : Type a, b : X u, v : a = b p : u = v q : u @ 1 = v @ 1 theta : whiskerR p 1 @ 1 = 1 @ q
(((rlucancel_inv (1 [-] theta))^ @ wlrnat 1 p) @
rlucancel_inv (theta [-] 1 )) @ concat_p1 q =
concat_1p q
Proof .X : Type a, b : X u, v : a = b p : u = v q : u @ 1 = v @ 1 theta : whiskerR p 1 @ 1 = 1 @ q
(((rlucancel_inv (1 [-] theta))^ @ wlrnat 1 p) @
rlucancel_inv (theta [-] 1 )) @ concat_p1 q =
concat_1p q
revert q theta.X : Type a, b : X u, v : a = b p : u = v
forall (q : u @ 1 = v @ 1 )
(theta : whiskerR p 1 @ 1 = 1 @ q),
(((rlucancel_inv (1 [-] theta))^ @ wlrnat 1 p) @
rlucancel_inv (theta [-] 1 )) @ concat_p1 q =
concat_1p q
snapply equiv_path_ind_rlucancel. X : Type a, b : X u, v : a = b p : u = v
(fun (q : u @ 1 = v @ 1 )
(s : whiskerR p 1 @ 1 = 1 @ q) =>
(((rlucancel_inv (1 [-] s))^ @ wlrnat 1 p) @
rlucancel_inv (s [-] 1 )) @ concat_p1 q = concat_1p q)
(whiskerR p 1 ) (rlucancel 1 )
by destruct p.
Defined .
Definition eh_1p {X } {a : X} (p : idpath a = idpath a) :
eh 1 p @ concat_p1 p = concat_1p p.X : Type a : X p : 1 = 1
eh 1 p @ concat_p1 p = concat_1p p
Proof .X : Type a : X p : 1 = 1
eh 1 p @ concat_p1 p = concat_1p p
exact (eh_1p_gen p (urnat p)).
Defined .
Local Definition eh_p1_gen {X } {a b : X} {u v : a = b} (p : u = v) {q } (theta : whiskerL 1 p @ 1 = 1 @ q) :
(rlucancel_inv (theta [-] 1 ))^ @ wlrnat p 1 @ rlucancel_inv (1 [-] theta) @ concat_1p q = concat_p1 q.X : Type a, b : X u, v : a = b p : u = v q : 1 @ u = 1 @ vtheta : whiskerL 1 p @ 1 = 1 @ q
(((rlucancel_inv (theta [-] 1 ))^ @ wlrnat p 1 ) @
rlucancel_inv (1 [-] theta)) @ concat_1p q =
concat_p1 q
Proof .X : Type a, b : X u, v : a = b p : u = v q : 1 @ u = 1 @ vtheta : whiskerL 1 p @ 1 = 1 @ q
(((rlucancel_inv (theta [-] 1 ))^ @ wlrnat p 1 ) @
rlucancel_inv (1 [-] theta)) @ concat_1p q =
concat_p1 q
revert q theta.X : Type a, b : X u, v : a = b p : u = v
forall (q : 1 @ u = 1 @ v)
(theta : whiskerL 1 p @ 1 = 1 @ q),
(((rlucancel_inv (theta [-] 1 ))^ @ wlrnat p 1 ) @
rlucancel_inv (1 [-] theta)) @ concat_1p q =
concat_p1 q
snapply equiv_path_ind_rlucancel. X : Type a, b : X u, v : a = b p : u = v
(fun (q : 1 @ u = 1 @ v)
(s : whiskerL 1 p @ 1 = 1 @ q) =>
(((rlucancel_inv (s [-] 1 ))^ @ wlrnat p 1 ) @
rlucancel_inv (1 [-] s)) @ concat_1p q = concat_p1 q)
(whiskerL 1 p) (rlucancel 1 )
by destruct p.
Defined .
Definition eh_p1 {X } {a : X} (p : idpath a = idpath a) :
eh p 1 @ concat_1p p = concat_p1 p.X : Type a : X p : 1 = 1
eh p 1 @ concat_1p p = concat_p1 p
Proof .X : Type a : X p : 1 = 1
eh p 1 @ concat_1p p = concat_p1 p
exact (eh_p1_gen p (ulnat p)).
Defined .
(** Naturality of Eckmann-Hilton. *)
Definition ehlnat {X } {a : X} (u : idpath a = idpath a) {x y } (p : x = y) :
whiskerL u p @ eh u y = eh u x @ whiskerR p u.X : Type a : X u, x, y : 1 = 1 p : x = y
whiskerL u p @ eh u y = eh u x @ whiskerR p u
Proof .X : Type a : X u, x, y : 1 = 1 p : x = y
whiskerL u p @ eh u y = eh u x @ whiskerR p u
destruct p.X : Type a : X u, x : 1 = 1
whiskerL u 1 @ eh u x = eh u x @ whiskerR 1 u
exact (lrucancel 1 ).
Defined .
Definition ehrnat {X } {a : X} {u v } (p : u = v) (x : idpath a = idpath a) :
whiskerR p x @ eh v x = eh u x @ whiskerL x p.X : Type a : X u, v : 1 = 1 p : u = v x : 1 = 1
whiskerR p x @ eh v x = eh u x @ whiskerL x p
Proof .X : Type a : X u, v : 1 = 1 p : u = v x : 1 = 1
whiskerR p x @ eh v x = eh u x @ whiskerL x p
destruct p.X : Type a : X u, x : 1 = 1
whiskerR 1 x @ eh u x = eh u x @ whiskerL x 1
exact (lrucancel 1 ).
Defined .
(** Naturality of Eckmann-Hilton when the fixed path is 1. *)
Definition ehlnat_1p {X } {a : X} {u v : idpath a = idpath a} (p : u = v) :
(ehlnat 1 p [I] urnat p) @ whiskerR (eh_1p u) _ = whiskerL _ (eh_1p v) @ ulnat p.X : Type a : X u, v : 1 = 1 p : u = v
(ehlnat 1 p [I] urnat p) @ whiskerR (eh_1p u) p =
whiskerL (whiskerL 1 p) (eh_1p v) @ ulnat p
Proof .X : Type a : X u, v : 1 = 1 p : u = v
(ehlnat 1 p [I] urnat p) @ whiskerR (eh_1p u) p =
whiskerL (whiskerL 1 p) (eh_1p v) @ ulnat p
destruct p.X : Type a : X u : 1 = 1
(ehlnat 1 1 [I] urnat 1 ) @ whiskerR (eh_1p u) 1 =
whiskerL (whiskerL 1 1 ) (eh_1p u) @ ulnat 1
apply lrucancel_sHs_1_pp.
Defined .
Definition ehrnat_p1 {X } {a : X} {u v : idpath a = idpath a} (p : u = v) :
(ehrnat p 1 [I] ulnat p) @ whiskerR (eh_p1 u) _ = whiskerL _ (eh_p1 v) @ urnat p.X : Type a : X u, v : 1 = 1 p : u = v
(ehrnat p 1 [I] ulnat p) @ whiskerR (eh_p1 u) p =
whiskerL (whiskerR p 1 ) (eh_p1 v) @ urnat p
Proof .X : Type a : X u, v : 1 = 1 p : u = v
(ehrnat p 1 [I] ulnat p) @ whiskerR (eh_p1 u) p =
whiskerL (whiskerR p 1 ) (eh_p1 v) @ urnat p
destruct p.X : Type a : X u : 1 = 1
(ehrnat 1 1 [I] ulnat 1 ) @ whiskerR (eh_p1 u) 1 =
whiskerL (whiskerR 1 1 ) (eh_p1 u) @ urnat 1
apply lrucancel_sHs_1_pp.
Defined .
(* These lemmas should probably be in the library in some form. *)
Local Definition concat_p_pp_pp_p {A } {u v x y : A} (p : u = v) (q : v = x) (r : x = y) :
concat_p_pp p q r @ concat_pp_p p q r = 1 .A : Type u, v, x, y : A p : u = v q : v = x r : x = y
concat_p_pp p q r @ concat_pp_p p q r = 1
Proof .A : Type u, v, x, y : A p : u = v q : v = x r : x = y
concat_p_pp p q r @ concat_pp_p p q r = 1
by destruct p, q, r.
Defined .
Local Definition concat_pp_p_p_pp {A } {u v x y : A} (p : u = v) (q : v = x) (r : x = y) :
concat_pp_p p q r @ concat_p_pp p q r = 1 .A : Type u, v, x, y : A p : u = v q : v = x r : x = y
concat_pp_p p q r @ concat_p_pp p q r = 1
Proof .A : Type u, v, x, y : A p : u = v q : v = x r : x = y
concat_pp_p p q r @ concat_p_pp p q r = 1
by destruct p, q, r.
Defined .
(* These lemmas are in the library but with worse computational behavior. *)
Local Definition whiskerL_pp {A } {a b c : A} (u : a = b) {v w z : b = c} (p : v = w) (q : w = z) :
whiskerL u (p @ q) = whiskerL u p @ whiskerL u q.A : Type a, b, c : A u : a = b v, w, z : b = c p : v = w q : w = z
whiskerL u (p @ q) = whiskerL u p @ whiskerL u q
Proof .A : Type a, b, c : A u : a = b v, w, z : b = c p : v = w q : w = z
whiskerL u (p @ q) = whiskerL u p @ whiskerL u q
by destruct p, q.
Defined .
Local Definition whiskerR_pp {A } {a b c : A} {u v w : a = b} (z : b = c) (p : u = v) (q : v = w) :
whiskerR (p @ q) z = whiskerR p z @ whiskerR q z.A : Type a, b, c : A u, v, w : a = b z : b = c p : u = v q : v = w
whiskerR (p @ q) z = whiskerR p z @ whiskerR q z
Proof .A : Type a, b, c : A u, v, w : a = b z : b = c p : u = v q : v = w
whiskerR (p @ q) z = whiskerR p z @ whiskerR q z
by destruct p, q.
Defined .
(** We now prove that [ulnat (p @ q)] suitably relates to [ulnat p] and [ulnat q]. *)
Definition ulnat_pp {X } {a b : X} {u v w : a = b} (p : u = v) (q : v = w) :
ulnat p [-] ulnat q = whiskerR (whiskerL_pp _ p q)^ _ @ ulnat (p @ q).X : Type a, b : X u, v, w : a = b p : u = v q : v = w
ulnat p [-] ulnat q =
whiskerR (whiskerL_pp 1 p q)^ (concat_1p w) @
ulnat (p @ q)
Proof .X : Type a, b : X u, v, w : a = b p : u = v q : v = w
ulnat p [-] ulnat q =
whiskerR (whiskerL_pp 1 p q)^ (concat_1p w) @
ulnat (p @ q)
by destruct p, q, u.
Defined .
(** We now prove that [urnat (p @ q)] suitably relates to [urnat p] and [urnat q]. *)
Definition urnat_pp {X } {a b : X} {u v w : a = b} (p : u = v) (q : v = w) :
urnat p [-] urnat q = whiskerR (whiskerR_pp _ p q)^ _ @ urnat (p @ q).X : Type a, b : X u, v, w : a = b p : u = v q : v = w
urnat p [-] urnat q =
whiskerR (whiskerR_pp 1 p q)^ (concat_p1 w) @
urnat (p @ q)
Proof .X : Type a, b : X u, v, w : a = b p : u = v q : v = w
urnat p [-] urnat q =
whiskerR (whiskerR_pp 1 p q)^ (concat_p1 w) @
urnat (p @ q)
by destruct p, q, u.
Defined .
(** We now prove that [ehlnat u (p @ q)] suitably relates to [ehlnat u p] and [ehlnat u q]. *)
Definition ehlnat_pp {X } {a : X} (u : idpath a = idpath a) {v w : idpath a = idpath a} (p : v = 1 ) (q : 1 = w) :
(ehlnat u p [-] ehlnat u q) @ whiskerL _ (whiskerR_pp _ p q)^ =
(whiskerR (whiskerL_pp _ p q)^ _) @ ehlnat u (p @ q).X : Type a : X u, v, w : 1 = 1 p : v = 1 q : 1 = w
(ehlnat u p [-] ehlnat u q) @
whiskerL (eh u v) (whiskerR_pp u p q)^ =
whiskerR (whiskerL_pp u p q)^ (eh u w) @
ehlnat u (p @ q)
Proof .X : Type a : X u, v, w : 1 = 1 p : v = 1 q : 1 = w
(ehlnat u p [-] ehlnat u q) @
whiskerL (eh u v) (whiskerR_pp u p q)^ =
whiskerR (whiskerL_pp u p q)^ (eh u w) @
ehlnat u (p @ q)
revert v p.X : Type a : X u, w : 1 = 1 q : 1 = w
forall (v : 1 = 1 ) (p : v = 1 ),
(ehlnat u p [-] ehlnat u q) @
whiskerL (eh u v) (whiskerR_pp u p q)^ =
whiskerR (whiskerL_pp u p q)^ (eh u w) @
ehlnat u (p @ q)
snapply (equiv_path_ind (equiv_path_inverse _)). X : Type a : X u, w : 1 = 1 q : 1 = w
(fun (b : 1 = 1 ) (x : (fun y : 1 = 1 => y = 1 ) b) =>
(ehlnat u x [-] ehlnat u q) @
whiskerL (eh u b) (whiskerR_pp u x q)^ =
whiskerR (whiskerL_pp u x q)^ (eh u w) @
ehlnat u (x @ q)) 1 (equiv_path_inverse 1 1 1 )
destruct q.X : Type a : X u : 1 = 1
(ehlnat u (equiv_path_inverse 1 1 1 ) [-] ehlnat u 1 ) @
whiskerL (eh u 1 )
(whiskerR_pp u (equiv_path_inverse 1 1 1 ) 1 )^ =
whiskerR (whiskerL_pp u (equiv_path_inverse 1 1 1 ) 1 )^
(eh u 1 ) @ ehlnat u (equiv_path_inverse 1 1 1 @ 1 )
apply rlucancel, lrucancel_sVs_1.
Defined .
(** We now prove that [ehrnat (p @ q) w] suitably relates to [ehrnat p w] and [ehrnat q w]. *)
Definition ehrnat_pp {X } {a : X} {u v : idpath a = idpath a} (p : u = 1 ) (q : 1 = v) (w : idpath a = idpath a) :
(ehrnat p w [-] ehrnat q w) @ whiskerL _ (whiskerL_pp _ p q)^ =
(whiskerR (whiskerR_pp _ p q)^ _) @ ehrnat (p @ q) w.X : Type a : X u, v : 1 = 1 p : u = 1 q : 1 = vw : 1 = 1
(ehrnat p w [-] ehrnat q w) @
whiskerL (eh u w) (whiskerL_pp w p q)^ =
whiskerR (whiskerR_pp w p q)^ (eh v w) @
ehrnat (p @ q) w
Proof .X : Type a : X u, v : 1 = 1 p : u = 1 q : 1 = vw : 1 = 1
(ehrnat p w [-] ehrnat q w) @
whiskerL (eh u w) (whiskerL_pp w p q)^ =
whiskerR (whiskerR_pp w p q)^ (eh v w) @
ehrnat (p @ q) w
revert u p.X : Type a : X v : 1 = 1 q : 1 = vw : 1 = 1
forall (u : 1 = 1 ) (p : u = 1 ),
(ehrnat p w [-] ehrnat q w) @
whiskerL (eh u w) (whiskerL_pp w p q)^ =
whiskerR (whiskerR_pp w p q)^ (eh v w) @
ehrnat (p @ q) w
snapply (equiv_path_ind (equiv_path_inverse _)). X : Type a : X v : 1 = 1 q : 1 = vw : 1 = 1
(fun (b : 1 = 1 ) (x : (fun y : 1 = 1 => y = 1 ) b) =>
(ehrnat x w [-] ehrnat q w) @
whiskerL (eh b w) (whiskerL_pp w x q)^ =
whiskerR (whiskerR_pp w x q)^ (eh v w) @
ehrnat (x @ q) w) 1 (equiv_path_inverse 1 1 1 )
destruct q.X : Type a : X w : 1 = 1
(ehrnat (equiv_path_inverse 1 1 1 ) w [-] ehrnat 1 w) @
whiskerL (eh 1 w)
(whiskerL_pp w (equiv_path_inverse 1 1 1 ) 1 )^ =
whiskerR (whiskerR_pp w (equiv_path_inverse 1 1 1 ) 1 )^
(eh 1 w) @ ehrnat (equiv_path_inverse 1 1 1 @ 1 ) w
cbn .X : Type a : X w : 1 = 1
((concat_1p (eh 1 w) @ 1 ) @ (concat_p1 (eh 1 w))^
[-] (concat_1p (eh 1 w) @ 1 ) @ (concat_p1 (eh 1 w))^) @
1 =
1 @ ((concat_1p (eh 1 w) @ 1 ) @ (concat_p1 (eh 1 w))^)
apply rlucancel, lrucancel_sVs_1.
Defined .
(** We now prove that [wlrnat p (q @ r)] suitably relates to [wlrnat p q] and [wlrnat q p]. *)
Definition wlrnat_p_pp {X } {a b c : X} {u v w : a = b} {x y : b = c} (p : x = y) (q : u = v) (r : v = w) :
(wlrnat p q [I] wlrnat p r) @ whiskerR (whiskerR_pp _ q r)^ _ =
whiskerL _ (whiskerR_pp _ q r)^ @ wlrnat p (q @ r).X : Type a, b, c : X u, v, w : a = b x, y : b = c p : x = y q : u = v r : v = w
(wlrnat p q [I] wlrnat p r) @
whiskerR (whiskerR_pp x q r)^ (whiskerL w p) =
whiskerL (whiskerL u p) (whiskerR_pp y q r)^ @
wlrnat p (q @ r)
Proof .X : Type a, b, c : X u, v, w : a = b x, y : b = c p : x = y q : u = v r : v = w
(wlrnat p q [I] wlrnat p r) @
whiskerR (whiskerR_pp x q r)^ (whiskerL w p) =
whiskerL (whiskerL u p) (whiskerR_pp y q r)^ @
wlrnat p (q @ r)
by destruct p, q, r.
Defined .
(** We now prove that [wlrnat (p @ q) r] suitably relates to [wlrnat p r] and [wlrnat q r]. *)
Definition wlrnat_pp_p {X } {a b c : X} {u v : a = b} {x y z : b = c} (p : x = y) (q : y = z) (r : u = v) :
(wlrnat p r [-] wlrnat q r) @ whiskerL _ (whiskerL_pp _ p q)^ =
whiskerR (whiskerL_pp _ p q)^ _ @ wlrnat (p @ q) r.X : Type a, b, c : X u, v : a = b x, y, z : b = c p : x = y q : y = z r : u = v
(wlrnat p r [-] wlrnat q r) @
whiskerL (whiskerR r x) (whiskerL_pp v p q)^ =
whiskerR (whiskerL_pp u p q)^ (whiskerR r z) @
wlrnat (p @ q) r
Proof .X : Type a, b, c : X u, v : a = b x, y, z : b = c p : x = y q : y = z r : u = v
(wlrnat p r [-] wlrnat q r) @
whiskerL (whiskerR r x) (whiskerL_pp v p q)^ =
whiskerR (whiskerL_pp u p q)^ (whiskerR r z) @
wlrnat (p @ q) r
by destruct p, q, r.
Defined .
(** We now prove that [wrlnat p q] suitably relates to [wlrnat q p]. *)
Definition wlrnat_V {X } {a : X} {u v x y : idpath a = idpath a} p q :
whiskerR (wlrnat p q) (eh v y) @ (ehrnat q x [-] ehlnat v p) =
(ehlnat u p [-] ehrnat q y) @ whiskerL (eh u x) (wlrnat q p)^.X : Type a : X u, v, x, y : 1 = 1 p : x = y q : u = v
whiskerR (wlrnat p q) (eh v y) @
(ehrnat q x [-] ehlnat v p) =
(ehlnat u p [-] ehrnat q y) @
whiskerL (eh u x) (wlrnat q p)^
Proof .X : Type a : X u, v, x, y : 1 = 1 p : x = y q : u = v
whiskerR (wlrnat p q) (eh v y) @
(ehrnat q x [-] ehlnat v p) =
(ehlnat u p [-] ehrnat q y) @
whiskerL (eh u x) (wlrnat q p)^
destruct p, q.X : Type a : X u, x : 1 = 1
whiskerR (wlrnat 1 1 ) (eh u x) @
(ehrnat 1 x [-] ehlnat u 1 ) =
(ehlnat u 1 [-] ehrnat 1 x) @
whiskerL (eh u x) (wlrnat 1 1 )^
exact (lrucancel 1 ).
Defined .
(** Coherence #1: We now prove that [eh p (q @ r)] suitably relates to [eh p q] and [eh p r]. *)
Section eh_p_pp .
Context {X : Type }.
(* 0-paths *)
Context {a b c d e f : X}.
(* 1-paths *)
Context {wlx0 x0 : a = b}.
Context {wlx1 x1 : c = d}.
Context {wlx2 x2 : e = f}.
Context {wry0 y0 : b = d}.
Context {wry1 y1 : a = c}.
Context {wrz0 z0 : d = f}.
Context {wrz1 z1 : c = e}.
Context {wryz0 : b = f}.
Context {wryz1 : a = e}.
(* 2-paths *)
Context {ulnat_x0 : wlx0 @ 1 = 1 @ x0}.
Context {ulnat_x1 : wlx1 @ 1 = 1 @ x1}.
Context {ulnat_x2 : wlx2 @ 1 = 1 @ x2}.
Context {urnat_y0 : wry0 @ 1 = 1 @ y0}.
Context {urnat_y1 : wry1 @ 1 = 1 @ y1}.
Context {urnat_z0 : wrz0 @ 1 = 1 @ z0}.
Context {urnat_z1 : wrz1 @ 1 = 1 @ z1}.
Context {urnat_yz0 : wryz0 @ 1 = 1 @ (y0 @ z0)}.
Context {urnat_yz1 : wryz1 @ 1 = 1 @ (y1 @ z1)}.
Context {wlrnat_x_y : wlx0 @ wry0 = wry1 @ wlx1}.
Context {wlrnat_x_z : wlx1 @ wrz0 = wrz1 @ wlx2}.
Context {wlrnat_x_yz : wlx0 @ wryz0 = wryz1 @ wlx2}.
Context {wrpp_yz0 : wry0 @ wrz0 = wryz0}.
Context {wrpp_yz1 : wry1 @ wrz1 = wryz1}.
(* 3-paths *)
Hypothesis H_urnat_yz0 :
(urnat_y0 [-] urnat_z0) = whiskerR wrpp_yz0 _ @ urnat_yz0.
Hypothesis H_urnat_yz1 :
(urnat_y1 [-] urnat_z1) = whiskerR wrpp_yz1 _ @ urnat_yz1.
Hypothesis H_wlrnat_x_yz :
(wlrnat_x_y [I] wlrnat_x_z) @ whiskerR wrpp_yz1 _ =
whiskerL _ wrpp_yz0 @ wlrnat_x_yz.
(* the coherence *)
Definition eh_p_pp_gen :
let EH_x_y := (rlucancel_inv (ulnat_x0 [-] urnat_y0))^ @
wlrnat_x_y @ rlucancel_inv (urnat_y1 [-] ulnat_x1) in
let EH_x_z := (rlucancel_inv (ulnat_x1 [-] urnat_z0))^ @
wlrnat_x_z @ rlucancel_inv (urnat_z1 [-] ulnat_x2) in
let EH_x_yz := (rlucancel_inv (ulnat_x0 [-] urnat_yz0))^ @
wlrnat_x_yz @ rlucancel_inv (urnat_yz1 [-] ulnat_x2) in
EH_x_yz @ (concat_pp_p _ _ _ @ whiskerL _ EH_x_z^) =
concat_p_pp _ _ _ @ whiskerR EH_x_y _ @ concat_pp_p _ _ _.X : Type a, b, c, d, e, f : X wlx0, x0 : a = b wlx1, x1 : c = d wlx2, x2 : e = f wry0, y0 : b = d wry1, y1 : a = c wrz0, z0 : d = f wrz1, z1 : c = e wryz0 : b = f wryz1 : a = e ulnat_x0 : wlx0 @ 1 = 1 @ x0 ulnat_x1 : wlx1 @ 1 = 1 @ x1 ulnat_x2 : wlx2 @ 1 = 1 @ x2 urnat_y0 : wry0 @ 1 = 1 @ y0 urnat_y1 : wry1 @ 1 = 1 @ y1 urnat_z0 : wrz0 @ 1 = 1 @ z0 urnat_z1 : wrz1 @ 1 = 1 @ z1 urnat_yz0 : wryz0 @ 1 = 1 @ (y0 @ z0) urnat_yz1 : wryz1 @ 1 = 1 @ (y1 @ z1) wlrnat_x_y : wlx0 @ wry0 = wry1 @ wlx1 wlrnat_x_z : wlx1 @ wrz0 = wrz1 @ wlx2 wlrnat_x_yz : wlx0 @ wryz0 = wryz1 @ wlx2 wrpp_yz0 : wry0 @ wrz0 = wryz0 wrpp_yz1 : wry1 @ wrz1 = wryz1 H_urnat_yz0 : urnat_y0 [-] urnat_z0 =
whiskerR wrpp_yz0 1 @ urnat_yz0 H_urnat_yz1 : urnat_y1 [-] urnat_z1 =
whiskerR wrpp_yz1 1 @ urnat_yz1 H_wlrnat_x_yz : (wlrnat_x_y [I] wlrnat_x_z) @
whiskerR wrpp_yz1 wlx2 =
whiskerL wlx0 wrpp_yz0 @ wlrnat_x_yz
let EH_x_y :=
((rlucancel_inv (ulnat_x0 [-] urnat_y0))^ @
wlrnat_x_y) @ rlucancel_inv (urnat_y1 [-] ulnat_x1)
in
let EH_x_z :=
((rlucancel_inv (ulnat_x1 [-] urnat_z0))^ @
wlrnat_x_z) @ rlucancel_inv (urnat_z1 [-] ulnat_x2)
in
let EH_x_yz :=
((rlucancel_inv (ulnat_x0 [-] urnat_yz0))^ @
wlrnat_x_yz) @
rlucancel_inv (urnat_yz1 [-] ulnat_x2) in
EH_x_yz @ (concat_pp_p y1 z1 x2 @ whiskerL y1 EH_x_z^) =
(concat_p_pp x0 y0 z0 @ whiskerR EH_x_y z0) @
concat_pp_p y1 x1 z0
Proof .X : Type a, b, c, d, e, f : X wlx0, x0 : a = b wlx1, x1 : c = d wlx2, x2 : e = f wry0, y0 : b = d wry1, y1 : a = c wrz0, z0 : d = f wrz1, z1 : c = e wryz0 : b = f wryz1 : a = e ulnat_x0 : wlx0 @ 1 = 1 @ x0 ulnat_x1 : wlx1 @ 1 = 1 @ x1 ulnat_x2 : wlx2 @ 1 = 1 @ x2 urnat_y0 : wry0 @ 1 = 1 @ y0 urnat_y1 : wry1 @ 1 = 1 @ y1 urnat_z0 : wrz0 @ 1 = 1 @ z0 urnat_z1 : wrz1 @ 1 = 1 @ z1 urnat_yz0 : wryz0 @ 1 = 1 @ (y0 @ z0) urnat_yz1 : wryz1 @ 1 = 1 @ (y1 @ z1) wlrnat_x_y : wlx0 @ wry0 = wry1 @ wlx1 wlrnat_x_z : wlx1 @ wrz0 = wrz1 @ wlx2 wlrnat_x_yz : wlx0 @ wryz0 = wryz1 @ wlx2 wrpp_yz0 : wry0 @ wrz0 = wryz0 wrpp_yz1 : wry1 @ wrz1 = wryz1 H_urnat_yz0 : urnat_y0 [-] urnat_z0 =
whiskerR wrpp_yz0 1 @ urnat_yz0 H_urnat_yz1 : urnat_y1 [-] urnat_z1 =
whiskerR wrpp_yz1 1 @ urnat_yz1 H_wlrnat_x_yz : (wlrnat_x_y [I] wlrnat_x_z) @
whiskerR wrpp_yz1 wlx2 =
whiskerL wlx0 wrpp_yz0 @ wlrnat_x_yz
let EH_x_y :=
((rlucancel_inv (ulnat_x0 [-] urnat_y0))^ @
wlrnat_x_y) @ rlucancel_inv (urnat_y1 [-] ulnat_x1)
in
let EH_x_z :=
((rlucancel_inv (ulnat_x1 [-] urnat_z0))^ @
wlrnat_x_z) @ rlucancel_inv (urnat_z1 [-] ulnat_x2)
in
let EH_x_yz :=
((rlucancel_inv (ulnat_x0 [-] urnat_yz0))^ @
wlrnat_x_yz) @
rlucancel_inv (urnat_yz1 [-] ulnat_x2) in
EH_x_yz @ (concat_pp_p y1 z1 x2 @ whiskerL y1 EH_x_z^) =
(concat_p_pp x0 y0 z0 @ whiskerR EH_x_y z0) @
concat_pp_p y1 x1 z0
apply moveR_Vp in H_urnat_yz0, H_urnat_yz1, H_wlrnat_x_yz.X : Type a, b, c, d, e, f : X wlx0, x0 : a = b wlx1, x1 : c = d wlx2, x2 : e = f wry0, y0 : b = d wry1, y1 : a = c wrz0, z0 : d = f wrz1, z1 : c = e wryz0 : b = f wryz1 : a = e ulnat_x0 : wlx0 @ 1 = 1 @ x0 ulnat_x1 : wlx1 @ 1 = 1 @ x1 ulnat_x2 : wlx2 @ 1 = 1 @ x2 urnat_y0 : wry0 @ 1 = 1 @ y0 urnat_y1 : wry1 @ 1 = 1 @ y1 urnat_z0 : wrz0 @ 1 = 1 @ z0 urnat_z1 : wrz1 @ 1 = 1 @ z1 urnat_yz0 : wryz0 @ 1 = 1 @ (y0 @ z0) urnat_yz1 : wryz1 @ 1 = 1 @ (y1 @ z1) wlrnat_x_y : wlx0 @ wry0 = wry1 @ wlx1 wlrnat_x_z : wlx1 @ wrz0 = wrz1 @ wlx2 wlrnat_x_yz : wlx0 @ wryz0 = wryz1 @ wlx2 wrpp_yz0 : wry0 @ wrz0 = wryz0 wrpp_yz1 : wry1 @ wrz1 = wryz1 H_urnat_yz0 : (whiskerR wrpp_yz0 1 )^ @
(urnat_y0 [-] urnat_z0) = urnat_yz0 H_urnat_yz1 : (whiskerR wrpp_yz1 1 )^ @
(urnat_y1 [-] urnat_z1) = urnat_yz1 H_wlrnat_x_yz : (whiskerL wlx0 wrpp_yz0)^ @
((wlrnat_x_y [I] wlrnat_x_z) @
whiskerR wrpp_yz1 wlx2) = wlrnat_x_yz
let EH_x_y :=
((rlucancel_inv (ulnat_x0 [-] urnat_y0))^ @
wlrnat_x_y) @ rlucancel_inv (urnat_y1 [-] ulnat_x1)
in
let EH_x_z :=
((rlucancel_inv (ulnat_x1 [-] urnat_z0))^ @
wlrnat_x_z) @ rlucancel_inv (urnat_z1 [-] ulnat_x2)
in
let EH_x_yz :=
((rlucancel_inv (ulnat_x0 [-] urnat_yz0))^ @
wlrnat_x_yz) @
rlucancel_inv (urnat_yz1 [-] ulnat_x2) in
EH_x_yz @ (concat_pp_p y1 z1 x2 @ whiskerL y1 EH_x_z^) =
(concat_p_pp x0 y0 z0 @ whiskerR EH_x_y z0) @
concat_pp_p y1 x1 z0
destruct H_urnat_yz0, H_urnat_yz1, H_wlrnat_x_yz.X : Type a, b, c, d, e, f : X wlx0, x0 : a = b wlx1, x1 : c = d wlx2, x2 : e = f wry0, y0 : b = d wry1, y1 : a = c wrz0, z0 : d = f wrz1, z1 : c = e wryz0 : b = f wryz1 : a = e ulnat_x0 : wlx0 @ 1 = 1 @ x0 ulnat_x1 : wlx1 @ 1 = 1 @ x1 ulnat_x2 : wlx2 @ 1 = 1 @ x2 urnat_y0 : wry0 @ 1 = 1 @ y0 urnat_y1 : wry1 @ 1 = 1 @ y1 urnat_z0 : wrz0 @ 1 = 1 @ z0 urnat_z1 : wrz1 @ 1 = 1 @ z1 wlrnat_x_y : wlx0 @ wry0 = wry1 @ wlx1 wlrnat_x_z : wlx1 @ wrz0 = wrz1 @ wlx2 wrpp_yz0 : wry0 @ wrz0 = wryz0 wrpp_yz1 : wry1 @ wrz1 = wryz1 H_urnat_yz0 : (whiskerR wrpp_yz0 1 )^ @
(urnat_y0 [-] urnat_z0) =
(whiskerR wrpp_yz0 1 )^ @
(urnat_y0 [-] urnat_z0) H_urnat_yz1 : (whiskerR wrpp_yz1 1 )^ @
(urnat_y1 [-] urnat_z1) =
(whiskerR wrpp_yz1 1 )^ @
(urnat_y1 [-] urnat_z1) H_wlrnat_x_yz : (whiskerL wlx0 wrpp_yz0)^ @
((wlrnat_x_y [I] wlrnat_x_z) @
whiskerR wrpp_yz1 wlx2) =
(whiskerL wlx0 wrpp_yz0)^ @
((wlrnat_x_y [I] wlrnat_x_z) @
whiskerR wrpp_yz1 wlx2)
let EH_x_y :=
((rlucancel_inv (ulnat_x0 [-] urnat_y0))^ @
wlrnat_x_y) @ rlucancel_inv (urnat_y1 [-] ulnat_x1)
in
let EH_x_z :=
((rlucancel_inv (ulnat_x1 [-] urnat_z0))^ @
wlrnat_x_z) @ rlucancel_inv (urnat_z1 [-] ulnat_x2)
in
let EH_x_yz :=
((rlucancel_inv
(ulnat_x0
[-] (whiskerR wrpp_yz0 1 )^ @
(urnat_y0 [-] urnat_z0)))^ @
((whiskerL wlx0 wrpp_yz0)^ @
((wlrnat_x_y [I] wlrnat_x_z) @
whiskerR wrpp_yz1 wlx2))) @
rlucancel_inv
((whiskerR wrpp_yz1 1 )^ @ (urnat_y1 [-] urnat_z1)
[-] ulnat_x2) in
EH_x_yz @ (concat_pp_p y1 z1 x2 @ whiskerL y1 EH_x_z^) =
(concat_p_pp x0 y0 z0 @ whiskerR EH_x_y z0) @
concat_pp_p y1 x1 z0
clear H_urnat_yz0 H_urnat_yz1 H_wlrnat_x_yz.X : Type a, b, c, d, e, f : X wlx0, x0 : a = b wlx1, x1 : c = d wlx2, x2 : e = f wry0, y0 : b = d wry1, y1 : a = c wrz0, z0 : d = f wrz1, z1 : c = e wryz0 : b = f wryz1 : a = e ulnat_x0 : wlx0 @ 1 = 1 @ x0 ulnat_x1 : wlx1 @ 1 = 1 @ x1 ulnat_x2 : wlx2 @ 1 = 1 @ x2 urnat_y0 : wry0 @ 1 = 1 @ y0 urnat_y1 : wry1 @ 1 = 1 @ y1 urnat_z0 : wrz0 @ 1 = 1 @ z0 urnat_z1 : wrz1 @ 1 = 1 @ z1 wlrnat_x_y : wlx0 @ wry0 = wry1 @ wlx1 wlrnat_x_z : wlx1 @ wrz0 = wrz1 @ wlx2 wrpp_yz0 : wry0 @ wrz0 = wryz0 wrpp_yz1 : wry1 @ wrz1 = wryz1
let EH_x_y :=
((rlucancel_inv (ulnat_x0 [-] urnat_y0))^ @
wlrnat_x_y) @ rlucancel_inv (urnat_y1 [-] ulnat_x1)
in
let EH_x_z :=
((rlucancel_inv (ulnat_x1 [-] urnat_z0))^ @
wlrnat_x_z) @ rlucancel_inv (urnat_z1 [-] ulnat_x2)
in
let EH_x_yz :=
((rlucancel_inv
(ulnat_x0
[-] (whiskerR wrpp_yz0 1 )^ @
(urnat_y0 [-] urnat_z0)))^ @
((whiskerL wlx0 wrpp_yz0)^ @
((wlrnat_x_y [I] wlrnat_x_z) @
whiskerR wrpp_yz1 wlx2))) @
rlucancel_inv
((whiskerR wrpp_yz1 1 )^ @ (urnat_y1 [-] urnat_z1)
[-] ulnat_x2) in
EH_x_yz @ (concat_pp_p y1 z1 x2 @ whiskerL y1 EH_x_z^) =
(concat_p_pp x0 y0 z0 @ whiskerR EH_x_y z0) @
concat_pp_p y1 x1 z0
destruct wrpp_yz0, wrpp_yz1.X : Type a, b, c, d, e, f : X wlx0, x0 : a = b wlx1, x1 : c = d wlx2, x2 : e = f wry0, y0 : b = d wry1, y1 : a = c wrz0, z0 : d = f wrz1, z1 : c = e ulnat_x0 : wlx0 @ 1 = 1 @ x0 ulnat_x1 : wlx1 @ 1 = 1 @ x1 ulnat_x2 : wlx2 @ 1 = 1 @ x2 urnat_y0 : wry0 @ 1 = 1 @ y0 urnat_y1 : wry1 @ 1 = 1 @ y1 urnat_z0 : wrz0 @ 1 = 1 @ z0 urnat_z1 : wrz1 @ 1 = 1 @ z1 wlrnat_x_y : wlx0 @ wry0 = wry1 @ wlx1 wlrnat_x_z : wlx1 @ wrz0 = wrz1 @ wlx2 wrpp_yz0 : wry0 @ wrz0 = wry0 @ wrz0 wrpp_yz1 : wry1 @ wrz1 = wry1 @ wrz1
let EH_x_y :=
((rlucancel_inv (ulnat_x0 [-] urnat_y0))^ @
wlrnat_x_y) @ rlucancel_inv (urnat_y1 [-] ulnat_x1)
in
let EH_x_z :=
((rlucancel_inv (ulnat_x1 [-] urnat_z0))^ @
wlrnat_x_z) @ rlucancel_inv (urnat_z1 [-] ulnat_x2)
in
let EH_x_yz :=
((rlucancel_inv
(ulnat_x0
[-] (whiskerR 1 1 )^ @ (urnat_y0 [-] urnat_z0)))^ @
((whiskerL wlx0 1 )^ @
((wlrnat_x_y [I] wlrnat_x_z) @ whiskerR 1 wlx2))) @
rlucancel_inv
((whiskerR 1 1 )^ @ (urnat_y1 [-] urnat_z1)
[-] ulnat_x2) in
EH_x_yz @ (concat_pp_p y1 z1 x2 @ whiskerL y1 EH_x_z^) =
(concat_p_pp x0 y0 z0 @ whiskerR EH_x_y z0) @
concat_pp_p y1 x1 z0
clear wrpp_yz0 wrpp_yz1.X : Type a, b, c, d, e, f : X wlx0, x0 : a = b wlx1, x1 : c = d wlx2, x2 : e = f wry0, y0 : b = d wry1, y1 : a = c wrz0, z0 : d = f wrz1, z1 : c = e ulnat_x0 : wlx0 @ 1 = 1 @ x0 ulnat_x1 : wlx1 @ 1 = 1 @ x1 ulnat_x2 : wlx2 @ 1 = 1 @ x2 urnat_y0 : wry0 @ 1 = 1 @ y0 urnat_y1 : wry1 @ 1 = 1 @ y1 urnat_z0 : wrz0 @ 1 = 1 @ z0 urnat_z1 : wrz1 @ 1 = 1 @ z1 wlrnat_x_y : wlx0 @ wry0 = wry1 @ wlx1 wlrnat_x_z : wlx1 @ wrz0 = wrz1 @ wlx2
let EH_x_y :=
((rlucancel_inv (ulnat_x0 [-] urnat_y0))^ @
wlrnat_x_y) @ rlucancel_inv (urnat_y1 [-] ulnat_x1)
in
let EH_x_z :=
((rlucancel_inv (ulnat_x1 [-] urnat_z0))^ @
wlrnat_x_z) @ rlucancel_inv (urnat_z1 [-] ulnat_x2)
in
let EH_x_yz :=
((rlucancel_inv
(ulnat_x0
[-] (whiskerR 1 1 )^ @ (urnat_y0 [-] urnat_z0)))^ @
((whiskerL wlx0 1 )^ @
((wlrnat_x_y [I] wlrnat_x_z) @ whiskerR 1 wlx2))) @
rlucancel_inv
((whiskerR 1 1 )^ @ (urnat_y1 [-] urnat_z1)
[-] ulnat_x2) in
EH_x_yz @ (concat_pp_p y1 z1 x2 @ whiskerL y1 EH_x_z^) =
(concat_p_pp x0 y0 z0 @ whiskerR EH_x_y z0) @
concat_pp_p y1 x1 z0
revert x0 ulnat_x0.X : Type a, b, c, d, e, f : X wlx0 : a = b wlx1, x1 : c = d wlx2, x2 : e = f wry0, y0 : b = d wry1, y1 : a = c wrz0, z0 : d = f wrz1, z1 : c = e ulnat_x1 : wlx1 @ 1 = 1 @ x1 ulnat_x2 : wlx2 @ 1 = 1 @ x2 urnat_y0 : wry0 @ 1 = 1 @ y0 urnat_y1 : wry1 @ 1 = 1 @ y1 urnat_z0 : wrz0 @ 1 = 1 @ z0 urnat_z1 : wrz1 @ 1 = 1 @ z1 wlrnat_x_y : wlx0 @ wry0 = wry1 @ wlx1 wlrnat_x_z : wlx1 @ wrz0 = wrz1 @ wlx2
forall (x0 : a = b) (ulnat_x0 : wlx0 @ 1 = 1 @ x0),
let EH_x_y :=
((rlucancel_inv (ulnat_x0 [-] urnat_y0))^ @
wlrnat_x_y) @ rlucancel_inv (urnat_y1 [-] ulnat_x1)
in
let EH_x_z :=
((rlucancel_inv (ulnat_x1 [-] urnat_z0))^ @
wlrnat_x_z) @ rlucancel_inv (urnat_z1 [-] ulnat_x2)
in
let EH_x_yz :=
((rlucancel_inv
(ulnat_x0
[-] (whiskerR 1 1 )^ @ (urnat_y0 [-] urnat_z0)))^ @
((whiskerL wlx0 1 )^ @
((wlrnat_x_y [I] wlrnat_x_z) @ whiskerR 1 wlx2))) @
rlucancel_inv
((whiskerR 1 1 )^ @ (urnat_y1 [-] urnat_z1)
[-] ulnat_x2) in
EH_x_yz @ (concat_pp_p y1 z1 x2 @ whiskerL y1 EH_x_z^) =
(concat_p_pp x0 y0 z0 @ whiskerR EH_x_y z0) @
concat_pp_p y1 x1 z0
snapply equiv_path_ind_rlucancel. X : Type a, b, c, d, e, f : X wlx0 : a = b wlx1, x1 : c = d wlx2, x2 : e = f wry0, y0 : b = d wry1, y1 : a = c wrz0, z0 : d = f wrz1, z1 : c = e ulnat_x1 : wlx1 @ 1 = 1 @ x1 ulnat_x2 : wlx2 @ 1 = 1 @ x2 urnat_y0 : wry0 @ 1 = 1 @ y0 urnat_y1 : wry1 @ 1 = 1 @ y1 urnat_z0 : wrz0 @ 1 = 1 @ z0 urnat_z1 : wrz1 @ 1 = 1 @ z1 wlrnat_x_y : wlx0 @ wry0 = wry1 @ wlx1 wlrnat_x_z : wlx1 @ wrz0 = wrz1 @ wlx2
(fun (q : a = b) (s : wlx0 @ 1 = 1 @ q) =>
let EH_x_y :=
((rlucancel_inv (s [-] urnat_y0))^ @ wlrnat_x_y) @
rlucancel_inv (urnat_y1 [-] ulnat_x1) in
let EH_x_z :=
((rlucancel_inv (ulnat_x1 [-] urnat_z0))^ @
wlrnat_x_z) @
rlucancel_inv (urnat_z1 [-] ulnat_x2) in
let EH_x_yz :=
((rlucancel_inv
(s
[-] (whiskerR 1 1 )^ @ (urnat_y0 [-] urnat_z0)))^ @
((whiskerL wlx0 1 )^ @
((wlrnat_x_y [I] wlrnat_x_z) @ whiskerR 1 wlx2))) @
rlucancel_inv
((whiskerR 1 1 )^ @ (urnat_y1 [-] urnat_z1)
[-] ulnat_x2) in
EH_x_yz @
(concat_pp_p y1 z1 x2 @ whiskerL y1 EH_x_z^) =
(concat_p_pp q y0 z0 @ whiskerR EH_x_y z0) @
concat_pp_p y1 x1 z0) wlx0 (rlucancel 1 )
revert x1 ulnat_x1.X : Type a, b, c, d, e, f : X wlx0 : a = b wlx1 : c = d wlx2, x2 : e = f wry0, y0 : b = d wry1, y1 : a = c wrz0, z0 : d = f wrz1, z1 : c = e ulnat_x2 : wlx2 @ 1 = 1 @ x2 urnat_y0 : wry0 @ 1 = 1 @ y0 urnat_y1 : wry1 @ 1 = 1 @ y1 urnat_z0 : wrz0 @ 1 = 1 @ z0 urnat_z1 : wrz1 @ 1 = 1 @ z1 wlrnat_x_y : wlx0 @ wry0 = wry1 @ wlx1 wlrnat_x_z : wlx1 @ wrz0 = wrz1 @ wlx2
forall (x1 : c = d) (ulnat_x1 : wlx1 @ 1 = 1 @ x1),
(fun (q : a = b) (s : wlx0 @ 1 = 1 @ q) =>
let EH_x_y :=
((rlucancel_inv (s [-] urnat_y0))^ @ wlrnat_x_y) @
rlucancel_inv (urnat_y1 [-] ulnat_x1) in
let EH_x_z :=
((rlucancel_inv (ulnat_x1 [-] urnat_z0))^ @
wlrnat_x_z) @
rlucancel_inv (urnat_z1 [-] ulnat_x2) in
let EH_x_yz :=
((rlucancel_inv
(s
[-] (whiskerR 1 1 )^ @ (urnat_y0 [-] urnat_z0)))^ @
((whiskerL wlx0 1 )^ @
((wlrnat_x_y [I] wlrnat_x_z) @ whiskerR 1 wlx2))) @
rlucancel_inv
((whiskerR 1 1 )^ @ (urnat_y1 [-] urnat_z1)
[-] ulnat_x2) in
EH_x_yz @
(concat_pp_p y1 z1 x2 @ whiskerL y1 EH_x_z^) =
(concat_p_pp q y0 z0 @ whiskerR EH_x_y z0) @
concat_pp_p y1 x1 z0) wlx0 (rlucancel 1 )
snapply equiv_path_ind_rlucancel. X : Type a, b, c, d, e, f : X wlx0 : a = b wlx1 : c = d wlx2, x2 : e = f wry0, y0 : b = d wry1, y1 : a = c wrz0, z0 : d = f wrz1, z1 : c = e ulnat_x2 : wlx2 @ 1 = 1 @ x2 urnat_y0 : wry0 @ 1 = 1 @ y0 urnat_y1 : wry1 @ 1 = 1 @ y1 urnat_z0 : wrz0 @ 1 = 1 @ z0 urnat_z1 : wrz1 @ 1 = 1 @ z1 wlrnat_x_y : wlx0 @ wry0 = wry1 @ wlx1 wlrnat_x_z : wlx1 @ wrz0 = wrz1 @ wlx2
(fun (q : c = d) (s : wlx1 @ 1 = 1 @ q) =>
(fun (q0 : a = b) (s0 : wlx0 @ 1 = 1 @ q0) =>
let EH_x_y :=
((rlucancel_inv (s0 [-] urnat_y0))^ @ wlrnat_x_y) @
rlucancel_inv (urnat_y1 [-] s) in
let EH_x_z :=
((rlucancel_inv (s [-] urnat_z0))^ @ wlrnat_x_z) @
rlucancel_inv (urnat_z1 [-] ulnat_x2) in
let EH_x_yz :=
((rlucancel_inv
(s0
[-] (whiskerR 1 1 )^ @ (urnat_y0 [-] urnat_z0)))^ @
((whiskerL wlx0 1 )^ @
((wlrnat_x_y [I] wlrnat_x_z) @ whiskerR 1 wlx2))) @
rlucancel_inv
((whiskerR 1 1 )^ @ (urnat_y1 [-] urnat_z1)
[-] ulnat_x2) in
EH_x_yz @
(concat_pp_p y1 z1 x2 @ whiskerL y1 EH_x_z^) =
(concat_p_pp q0 y0 z0 @ whiskerR EH_x_y z0) @
concat_pp_p y1 q z0) wlx0 (rlucancel 1 )) wlx1
(rlucancel 1 )
revert x2 ulnat_x2.X : Type a, b, c, d, e, f : X wlx0 : a = b wlx1 : c = d wlx2 : e = f wry0, y0 : b = d wry1, y1 : a = c wrz0, z0 : d = f wrz1, z1 : c = e urnat_y0 : wry0 @ 1 = 1 @ y0 urnat_y1 : wry1 @ 1 = 1 @ y1 urnat_z0 : wrz0 @ 1 = 1 @ z0 urnat_z1 : wrz1 @ 1 = 1 @ z1 wlrnat_x_y : wlx0 @ wry0 = wry1 @ wlx1 wlrnat_x_z : wlx1 @ wrz0 = wrz1 @ wlx2
forall (x2 : e = f) (ulnat_x2 : wlx2 @ 1 = 1 @ x2),
(fun (q : c = d) (s : wlx1 @ 1 = 1 @ q) =>
(fun (q0 : a = b) (s0 : wlx0 @ 1 = 1 @ q0) =>
let EH_x_y :=
((rlucancel_inv (s0 [-] urnat_y0))^ @ wlrnat_x_y) @
rlucancel_inv (urnat_y1 [-] s) in
let EH_x_z :=
((rlucancel_inv (s [-] urnat_z0))^ @ wlrnat_x_z) @
rlucancel_inv (urnat_z1 [-] ulnat_x2) in
let EH_x_yz :=
((rlucancel_inv
(s0
[-] (whiskerR 1 1 )^ @ (urnat_y0 [-] urnat_z0)))^ @
((whiskerL wlx0 1 )^ @
((wlrnat_x_y [I] wlrnat_x_z) @ whiskerR 1 wlx2))) @
rlucancel_inv
((whiskerR 1 1 )^ @ (urnat_y1 [-] urnat_z1)
[-] ulnat_x2) in
EH_x_yz @
(concat_pp_p y1 z1 x2 @ whiskerL y1 EH_x_z^) =
(concat_p_pp q0 y0 z0 @ whiskerR EH_x_y z0) @
concat_pp_p y1 q z0) wlx0 (rlucancel 1 )) wlx1
(rlucancel 1 )
snapply equiv_path_ind_rlucancel. X : Type a, b, c, d, e, f : X wlx0 : a = b wlx1 : c = d wlx2 : e = f wry0, y0 : b = d wry1, y1 : a = c wrz0, z0 : d = f wrz1, z1 : c = e urnat_y0 : wry0 @ 1 = 1 @ y0 urnat_y1 : wry1 @ 1 = 1 @ y1 urnat_z0 : wrz0 @ 1 = 1 @ z0 urnat_z1 : wrz1 @ 1 = 1 @ z1 wlrnat_x_y : wlx0 @ wry0 = wry1 @ wlx1 wlrnat_x_z : wlx1 @ wrz0 = wrz1 @ wlx2
(fun (q : e = f) (s : wlx2 @ 1 = 1 @ q) =>
(fun (q0 : c = d) (s0 : wlx1 @ 1 = 1 @ q0) =>
(fun (q1 : a = b) (s1 : wlx0 @ 1 = 1 @ q1) =>
let EH_x_y :=
((rlucancel_inv (s1 [-] urnat_y0))^ @ wlrnat_x_y) @
rlucancel_inv (urnat_y1 [-] s0) in
let EH_x_z :=
((rlucancel_inv (s0 [-] urnat_z0))^ @ wlrnat_x_z) @
rlucancel_inv (urnat_z1 [-] s) in
let EH_x_yz :=
((rlucancel_inv
(s1
[-] (whiskerR 1 1 )^ @
(urnat_y0 [-] urnat_z0)))^ @
((whiskerL wlx0 1 )^ @
((wlrnat_x_y [I] wlrnat_x_z) @ whiskerR 1 wlx2))) @
rlucancel_inv
((whiskerR 1 1 )^ @ (urnat_y1 [-] urnat_z1)
[-] s) in
EH_x_yz @
(concat_pp_p y1 z1 q @ whiskerL y1 EH_x_z^) =
(concat_p_pp q1 y0 z0 @ whiskerR EH_x_y z0) @
concat_pp_p y1 q0 z0) wlx0 (rlucancel 1 )) wlx1
(rlucancel 1 )) wlx2 (rlucancel 1 )
revert y0 urnat_y0.X : Type a, b, c, d, e, f : X wlx0 : a = b wlx1 : c = d wlx2 : e = f wry0 : b = d wry1, y1 : a = c wrz0, z0 : d = f wrz1, z1 : c = e urnat_y1 : wry1 @ 1 = 1 @ y1 urnat_z0 : wrz0 @ 1 = 1 @ z0 urnat_z1 : wrz1 @ 1 = 1 @ z1 wlrnat_x_y : wlx0 @ wry0 = wry1 @ wlx1 wlrnat_x_z : wlx1 @ wrz0 = wrz1 @ wlx2
forall (y0 : b = d) (urnat_y0 : wry0 @ 1 = 1 @ y0),
(fun (q : e = f) (s : wlx2 @ 1 = 1 @ q) =>
(fun (q0 : c = d) (s0 : wlx1 @ 1 = 1 @ q0) =>
(fun (q1 : a = b) (s1 : wlx0 @ 1 = 1 @ q1) =>
let EH_x_y :=
((rlucancel_inv (s1 [-] urnat_y0))^ @ wlrnat_x_y) @
rlucancel_inv (urnat_y1 [-] s0) in
let EH_x_z :=
((rlucancel_inv (s0 [-] urnat_z0))^ @ wlrnat_x_z) @
rlucancel_inv (urnat_z1 [-] s) in
let EH_x_yz :=
((rlucancel_inv
(s1
[-] (whiskerR 1 1 )^ @
(urnat_y0 [-] urnat_z0)))^ @
((whiskerL wlx0 1 )^ @
((wlrnat_x_y [I] wlrnat_x_z) @ whiskerR 1 wlx2))) @
rlucancel_inv
((whiskerR 1 1 )^ @ (urnat_y1 [-] urnat_z1)
[-] s) in
EH_x_yz @
(concat_pp_p y1 z1 q @ whiskerL y1 EH_x_z^) =
(concat_p_pp q1 y0 z0 @ whiskerR EH_x_y z0) @
concat_pp_p y1 q0 z0) wlx0 (rlucancel 1 )) wlx1
(rlucancel 1 )) wlx2 (rlucancel 1 )
snapply equiv_path_ind_rlucancel. X : Type a, b, c, d, e, f : X wlx0 : a = b wlx1 : c = d wlx2 : e = f wry0 : b = d wry1, y1 : a = c wrz0, z0 : d = f wrz1, z1 : c = e urnat_y1 : wry1 @ 1 = 1 @ y1 urnat_z0 : wrz0 @ 1 = 1 @ z0 urnat_z1 : wrz1 @ 1 = 1 @ z1 wlrnat_x_y : wlx0 @ wry0 = wry1 @ wlx1 wlrnat_x_z : wlx1 @ wrz0 = wrz1 @ wlx2
(fun (q : b = d) (s : wry0 @ 1 = 1 @ q) =>
(fun (q0 : e = f) (s0 : wlx2 @ 1 = 1 @ q0) =>
(fun (q1 : c = d) (s1 : wlx1 @ 1 = 1 @ q1) =>
(fun (q2 : a = b) (s2 : wlx0 @ 1 = 1 @ q2) =>
let EH_x_y :=
((rlucancel_inv (s2 [-] s))^ @ wlrnat_x_y) @
rlucancel_inv (urnat_y1 [-] s1) in
let EH_x_z :=
((rlucancel_inv (s1 [-] urnat_z0))^ @ wlrnat_x_z) @
rlucancel_inv (urnat_z1 [-] s0) in
let EH_x_yz :=
((rlucancel_inv
(s2 [-] (whiskerR 1 1 )^ @ (s [-] urnat_z0)))^ @
((whiskerL wlx0 1 )^ @
((wlrnat_x_y [I] wlrnat_x_z) @ whiskerR 1 wlx2))) @
rlucancel_inv
((whiskerR 1 1 )^ @ (urnat_y1 [-] urnat_z1)
[-] s0) in
EH_x_yz @
(concat_pp_p y1 z1 q0 @ whiskerL y1 EH_x_z^) =
(concat_p_pp q2 q z0 @ whiskerR EH_x_y z0) @
concat_pp_p y1 q1 z0) wlx0 (rlucancel 1 )) wlx1
(rlucancel 1 )) wlx2 (rlucancel 1 )) wry0
(rlucancel 1 )
revert y1 urnat_y1.X : Type a, b, c, d, e, f : X wlx0 : a = b wlx1 : c = d wlx2 : e = f wry0 : b = d wry1 : a = c wrz0, z0 : d = f wrz1, z1 : c = e urnat_z0 : wrz0 @ 1 = 1 @ z0 urnat_z1 : wrz1 @ 1 = 1 @ z1 wlrnat_x_y : wlx0 @ wry0 = wry1 @ wlx1 wlrnat_x_z : wlx1 @ wrz0 = wrz1 @ wlx2
forall (y1 : a = c) (urnat_y1 : wry1 @ 1 = 1 @ y1),
(fun (q : b = d) (s : wry0 @ 1 = 1 @ q) =>
(fun (q0 : e = f) (s0 : wlx2 @ 1 = 1 @ q0) =>
(fun (q1 : c = d) (s1 : wlx1 @ 1 = 1 @ q1) =>
(fun (q2 : a = b) (s2 : wlx0 @ 1 = 1 @ q2) =>
let EH_x_y :=
((rlucancel_inv (s2 [-] s))^ @ wlrnat_x_y) @
rlucancel_inv (urnat_y1 [-] s1) in
let EH_x_z :=
((rlucancel_inv (s1 [-] urnat_z0))^ @ wlrnat_x_z) @
rlucancel_inv (urnat_z1 [-] s0) in
let EH_x_yz :=
((rlucancel_inv
(s2 [-] (whiskerR 1 1 )^ @ (s [-] urnat_z0)))^ @
((whiskerL wlx0 1 )^ @
((wlrnat_x_y [I] wlrnat_x_z) @ whiskerR 1 wlx2))) @
rlucancel_inv
((whiskerR 1 1 )^ @ (urnat_y1 [-] urnat_z1)
[-] s0) in
EH_x_yz @
(concat_pp_p y1 z1 q0 @ whiskerL y1 EH_x_z^) =
(concat_p_pp q2 q z0 @ whiskerR EH_x_y z0) @
concat_pp_p y1 q1 z0) wlx0 (rlucancel 1 )) wlx1
(rlucancel 1 )) wlx2 (rlucancel 1 )) wry0
(rlucancel 1 )
snapply equiv_path_ind_rlucancel. X : Type a, b, c, d, e, f : X wlx0 : a = b wlx1 : c = d wlx2 : e = f wry0 : b = d wry1 : a = c wrz0, z0 : d = f wrz1, z1 : c = e urnat_z0 : wrz0 @ 1 = 1 @ z0 urnat_z1 : wrz1 @ 1 = 1 @ z1 wlrnat_x_y : wlx0 @ wry0 = wry1 @ wlx1 wlrnat_x_z : wlx1 @ wrz0 = wrz1 @ wlx2
(fun (q : a = c) (s : wry1 @ 1 = 1 @ q) =>
(fun (q0 : b = d) (s0 : wry0 @ 1 = 1 @ q0) =>
(fun (q1 : e = f) (s1 : wlx2 @ 1 = 1 @ q1) =>
(fun (q2 : c = d) (s2 : wlx1 @ 1 = 1 @ q2) =>
(fun (q3 : a = b) (s3 : wlx0 @ 1 = 1 @ q3) =>
let EH_x_y :=
((rlucancel_inv (s3 [-] s0))^ @ wlrnat_x_y) @
rlucancel_inv (s [-] s2) in
let EH_x_z :=
((rlucancel_inv (s2 [-] urnat_z0))^ @
wlrnat_x_z) @ rlucancel_inv (urnat_z1 [-] s1)
in
let EH_x_yz :=
((rlucancel_inv (s3 [-] ...^ @ ...))^ @
((whiskerL wlx0 1 )^ @
((...) @ whiskerR 1 wlx2))) @
rlucancel_inv
((whiskerR 1 1 )^ @ (s [-] urnat_z1) [-] s1)
in
EH_x_yz @
(concat_pp_p q z1 q1 @ whiskerL q EH_x_z^) =
(concat_p_pp q3 q0 z0 @ whiskerR EH_x_y z0) @
concat_pp_p q q2 z0) wlx0 (rlucancel 1 )) wlx1
(rlucancel 1 )) wlx2 (rlucancel 1 )) wry0
(rlucancel 1 )) wry1 (rlucancel 1 )
revert z0 urnat_z0.X : Type a, b, c, d, e, f : X wlx0 : a = b wlx1 : c = d wlx2 : e = f wry0 : b = d wry1 : a = c wrz0 : d = f wrz1, z1 : c = e urnat_z1 : wrz1 @ 1 = 1 @ z1 wlrnat_x_y : wlx0 @ wry0 = wry1 @ wlx1 wlrnat_x_z : wlx1 @ wrz0 = wrz1 @ wlx2
forall (z0 : d = f) (urnat_z0 : wrz0 @ 1 = 1 @ z0),
(fun (q : a = c) (s : wry1 @ 1 = 1 @ q) =>
(fun (q0 : b = d) (s0 : wry0 @ 1 = 1 @ q0) =>
(fun (q1 : e = f) (s1 : wlx2 @ 1 = 1 @ q1) =>
(fun (q2 : c = d) (s2 : wlx1 @ 1 = 1 @ q2) =>
(fun (q3 : a = b) (s3 : wlx0 @ 1 = 1 @ q3) =>
let EH_x_y :=
((rlucancel_inv (s3 [-] s0))^ @ wlrnat_x_y) @
rlucancel_inv (s [-] s2) in
let EH_x_z :=
((rlucancel_inv (s2 [-] urnat_z0))^ @
wlrnat_x_z) @ rlucancel_inv (urnat_z1 [-] s1)
in
let EH_x_yz :=
((rlucancel_inv (s3 [-] ...))^ @
((whiskerL wlx0 1 )^ @ (... @ ...))) @
rlucancel_inv
((whiskerR 1 1 )^ @ (s [-] urnat_z1) [-] s1)
in
EH_x_yz @
(concat_pp_p q z1 q1 @ whiskerL q EH_x_z^) =
(concat_p_pp q3 q0 z0 @ whiskerR EH_x_y z0) @
concat_pp_p q q2 z0) wlx0 (rlucancel 1 )) wlx1
(rlucancel 1 )) wlx2 (rlucancel 1 )) wry0
(rlucancel 1 )) wry1 (rlucancel 1 )
snapply equiv_path_ind_rlucancel. X : Type a, b, c, d, e, f : X wlx0 : a = b wlx1 : c = d wlx2 : e = f wry0 : b = d wry1 : a = c wrz0 : d = f wrz1, z1 : c = e urnat_z1 : wrz1 @ 1 = 1 @ z1 wlrnat_x_y : wlx0 @ wry0 = wry1 @ wlx1 wlrnat_x_z : wlx1 @ wrz0 = wrz1 @ wlx2
(fun (q : d = f) (s : wrz0 @ 1 = 1 @ q) =>
(fun (q0 : a = c) (s0 : wry1 @ 1 = 1 @ q0) =>
(fun (q1 : b = d) (s1 : wry0 @ 1 = 1 @ q1) =>
(fun (q2 : e = f) (s2 : wlx2 @ 1 = 1 @ q2) =>
(fun (q3 : c = d) (s3 : wlx1 @ 1 = 1 @ q3) =>
(fun (q4 : a = b) (s4 : wlx0 @ 1 = 1 @ q4) =>
let EH_x_y :=
((rlucancel_inv (s4 [-] s1))^ @ wlrnat_x_y) @
rlucancel_inv (s0 [-] s3) in
let EH_x_z :=
((rlucancel_inv (...))^ @ wlrnat_x_z) @
rlucancel_inv (urnat_z1 [-] s2) in
let EH_x_yz :=
((rlucancel_inv ...)^ @ (...^ @ ...)) @
rlucancel_inv (...^ @ ... [-] s2) in
EH_x_yz @
(concat_pp_p q0 z1 q2 @ whiskerL q0 EH_x_z^) =
(concat_p_pp q4 q1 q @ whiskerR EH_x_y q) @
concat_pp_p q0 q3 q) wlx0 (rlucancel 1 )) wlx1
(rlucancel 1 )) wlx2 (rlucancel 1 )) wry0
(rlucancel 1 )) wry1 (rlucancel 1 )) wrz0
(rlucancel 1 )
revert z1 urnat_z1.X : Type a, b, c, d, e, f : X wlx0 : a = b wlx1 : c = d wlx2 : e = f wry0 : b = d wry1 : a = c wrz0 : d = f wrz1 : c = e wlrnat_x_y : wlx0 @ wry0 = wry1 @ wlx1 wlrnat_x_z : wlx1 @ wrz0 = wrz1 @ wlx2
forall (z1 : c = e) (urnat_z1 : wrz1 @ 1 = 1 @ z1),
(fun (q : d = f) (s : wrz0 @ 1 = 1 @ q) =>
(fun (q0 : a = c) (s0 : wry1 @ 1 = 1 @ q0) =>
(fun (q1 : b = d) (s1 : wry0 @ 1 = 1 @ q1) =>
(fun (q2 : e = f) (s2 : wlx2 @ 1 = 1 @ q2) =>
(fun (q3 : c = d) (s3 : wlx1 @ 1 = 1 @ q3) =>
(fun (q4 : a = b) (s4 : wlx0 @ 1 = 1 @ q4) =>
let EH_x_y :=
((rlucancel_inv (...))^ @ wlrnat_x_y) @
rlucancel_inv (s0 [-] s3) in
let EH_x_z :=
((rlucancel_inv ...)^ @ wlrnat_x_z) @
rlucancel_inv (urnat_z1 [-] s2) in
let EH_x_yz :=
((...)^ @ (...)) @ rlucancel_inv (... [-] s2)
in
EH_x_yz @
(concat_pp_p q0 z1 q2 @ whiskerL q0 EH_x_z^) =
(concat_p_pp q4 q1 q @ whiskerR EH_x_y q) @
concat_pp_p q0 q3 q) wlx0 (rlucancel 1 )) wlx1
(rlucancel 1 )) wlx2 (rlucancel 1 )) wry0
(rlucancel 1 )) wry1 (rlucancel 1 )) wrz0
(rlucancel 1 )
snapply equiv_path_ind_rlucancel. X : Type a, b, c, d, e, f : X wlx0 : a = b wlx1 : c = d wlx2 : e = f wry0 : b = d wry1 : a = c wrz0 : d = f wrz1 : c = e wlrnat_x_y : wlx0 @ wry0 = wry1 @ wlx1 wlrnat_x_z : wlx1 @ wrz0 = wrz1 @ wlx2
(fun (q : c = e) (s : wrz1 @ 1 = 1 @ q) =>
(fun (q0 : d = f) (s0 : wrz0 @ 1 = 1 @ q0) =>
(fun (q1 : a = c) (s1 : wry1 @ 1 = 1 @ q1) =>
(fun (q2 : b = d) (s2 : wry0 @ 1 = 1 @ q2) =>
(fun (q3 : e = f) (s3 : wlx2 @ 1 = 1 @ q3) =>
(fun (q4 : c = d) (s4 : wlx1 @ 1 = 1 @ q4) =>
(fun (q5 : a = b) (s5 : wlx0 @ 1 = 1 @ q5) =>
let EH_x_y :=
((...)^ @ wlrnat_x_y) @
rlucancel_inv (s1 [-] s4) in
let EH_x_z :=
(...^ @ wlrnat_x_z) @ rlucancel_inv (...) in
let EH_x_yz := (...) @ rlucancel_inv ... in
EH_x_yz @ (...) = (...) @ concat_pp_p q1 q4 q0)
wlx0 (rlucancel 1 )) wlx1 (rlucancel 1 )) wlx2
(rlucancel 1 )) wry0 (rlucancel 1 )) wry1
(rlucancel 1 )) wrz0 (rlucancel 1 )) wrz1
(rlucancel 1 )
destruct wry0, wry1, wrz0, wrz1.X : Type a, b : X wlx0, wlx1, wlx2 : a = b wry0 : b = b wry1 : a = a wrz0 : b = b wrz1 : a = a wlrnat_x_y : wlx0 @ 1 = 1 @ wlx1 wlrnat_x_z : wlx1 @ 1 = 1 @ wlx2
let EH_x_y :=
((rlucancel_inv (rlucancel 1 [-] rlucancel 1 ))^ @
wlrnat_x_y) @
rlucancel_inv (rlucancel 1 [-] rlucancel 1 ) in
let EH_x_z :=
((rlucancel_inv (rlucancel 1 [-] rlucancel 1 ))^ @
wlrnat_x_z) @
rlucancel_inv (rlucancel 1 [-] rlucancel 1 ) in
let EH_x_yz :=
((rlucancel_inv
(rlucancel 1
[-] (whiskerR 1 1 )^ @
(rlucancel 1 [-] rlucancel 1 )))^ @
((whiskerL wlx0 1 )^ @
((wlrnat_x_y [I] wlrnat_x_z) @ whiskerR 1 wlx2))) @
rlucancel_inv
((whiskerR 1 1 )^ @ (rlucancel 1 [-] rlucancel 1 )
[-] rlucancel 1 ) in
EH_x_yz @ (concat_pp_p 1 1 wlx2 @ whiskerL 1 EH_x_z^) =
(concat_p_pp wlx0 1 1 @ whiskerR EH_x_y 1 ) @
concat_pp_p 1 wlx1 1
clear wry0 wry1 wrz0 wrz1.X : Type a, b : X wlx0, wlx1, wlx2 : a = b wlrnat_x_y : wlx0 @ 1 = 1 @ wlx1 wlrnat_x_z : wlx1 @ 1 = 1 @ wlx2
let EH_x_y :=
((rlucancel_inv (rlucancel 1 [-] rlucancel 1 ))^ @
wlrnat_x_y) @
rlucancel_inv (rlucancel 1 [-] rlucancel 1 ) in
let EH_x_z :=
((rlucancel_inv (rlucancel 1 [-] rlucancel 1 ))^ @
wlrnat_x_z) @
rlucancel_inv (rlucancel 1 [-] rlucancel 1 ) in
let EH_x_yz :=
((rlucancel_inv
(rlucancel 1
[-] (whiskerR 1 1 )^ @
(rlucancel 1 [-] rlucancel 1 )))^ @
((whiskerL wlx0 1 )^ @
((wlrnat_x_y [I] wlrnat_x_z) @ whiskerR 1 wlx2))) @
rlucancel_inv
((whiskerR 1 1 )^ @ (rlucancel 1 [-] rlucancel 1 )
[-] rlucancel 1 ) in
EH_x_yz @ (concat_pp_p 1 1 wlx2 @ whiskerL 1 EH_x_z^) =
(concat_p_pp wlx0 1 1 @ whiskerR EH_x_y 1 ) @
concat_pp_p 1 wlx1 1
revert wlx2 wlrnat_x_z.X : Type a, b : X wlx0, wlx1 : a = b wlrnat_x_y : wlx0 @ 1 = 1 @ wlx1
forall (wlx2 : a = b)
(wlrnat_x_z : wlx1 @ 1 = 1 @ wlx2),
let EH_x_y :=
((rlucancel_inv (rlucancel 1 [-] rlucancel 1 ))^ @
wlrnat_x_y) @
rlucancel_inv (rlucancel 1 [-] rlucancel 1 ) in
let EH_x_z :=
((rlucancel_inv (rlucancel 1 [-] rlucancel 1 ))^ @
wlrnat_x_z) @
rlucancel_inv (rlucancel 1 [-] rlucancel 1 ) in
let EH_x_yz :=
((rlucancel_inv
(rlucancel 1
[-] (whiskerR 1 1 )^ @
(rlucancel 1 [-] rlucancel 1 )))^ @
((whiskerL wlx0 1 )^ @
((wlrnat_x_y [I] wlrnat_x_z) @ whiskerR 1 wlx2))) @
rlucancel_inv
((whiskerR 1 1 )^ @ (rlucancel 1 [-] rlucancel 1 )
[-] rlucancel 1 ) in
EH_x_yz @ (concat_pp_p 1 1 wlx2 @ whiskerL 1 EH_x_z^) =
(concat_p_pp wlx0 1 1 @ whiskerR EH_x_y 1 ) @
concat_pp_p 1 wlx1 1
snapply equiv_path_ind_rlucancel. X : Type a, b : X wlx0, wlx1 : a = b wlrnat_x_y : wlx0 @ 1 = 1 @ wlx1
(fun (q : a = b) (s : wlx1 @ 1 = 1 @ q) =>
let EH_x_y :=
((rlucancel_inv (rlucancel 1 [-] rlucancel 1 ))^ @
wlrnat_x_y) @
rlucancel_inv (rlucancel 1 [-] rlucancel 1 ) in
let EH_x_z :=
((rlucancel_inv (rlucancel 1 [-] rlucancel 1 ))^ @ s) @
rlucancel_inv (rlucancel 1 [-] rlucancel 1 ) in
let EH_x_yz :=
((rlucancel_inv
(rlucancel 1
[-] (whiskerR 1 1 )^ @
(rlucancel 1 [-] rlucancel 1 )))^ @
((whiskerL wlx0 1 )^ @
((wlrnat_x_y [I] s) @ whiskerR 1 q))) @
rlucancel_inv
((whiskerR 1 1 )^ @ (rlucancel 1 [-] rlucancel 1 )
[-] rlucancel 1 ) in
EH_x_yz @ (concat_pp_p 1 1 q @ whiskerL 1 EH_x_z^) =
(concat_p_pp wlx0 1 1 @ whiskerR EH_x_y 1 ) @
concat_pp_p 1 wlx1 1 ) wlx1 (rlucancel 1 )
revert wlx1 wlrnat_x_y.X : Type a, b : X wlx0 : a = b
forall (wlx1 : a = b)
(wlrnat_x_y : wlx0 @ 1 = 1 @ wlx1),
(fun (q : a = b) (s : wlx1 @ 1 = 1 @ q) =>
let EH_x_y :=
((rlucancel_inv (rlucancel 1 [-] rlucancel 1 ))^ @
wlrnat_x_y) @
rlucancel_inv (rlucancel 1 [-] rlucancel 1 ) in
let EH_x_z :=
((rlucancel_inv (rlucancel 1 [-] rlucancel 1 ))^ @ s) @
rlucancel_inv (rlucancel 1 [-] rlucancel 1 ) in
let EH_x_yz :=
((rlucancel_inv
(rlucancel 1
[-] (whiskerR 1 1 )^ @
(rlucancel 1 [-] rlucancel 1 )))^ @
((whiskerL wlx0 1 )^ @
((wlrnat_x_y [I] s) @ whiskerR 1 q))) @
rlucancel_inv
((whiskerR 1 1 )^ @ (rlucancel 1 [-] rlucancel 1 )
[-] rlucancel 1 ) in
EH_x_yz @ (concat_pp_p 1 1 q @ whiskerL 1 EH_x_z^) =
(concat_p_pp wlx0 1 1 @ whiskerR EH_x_y 1 ) @
concat_pp_p 1 wlx1 1 ) wlx1 (rlucancel 1 )
snapply equiv_path_ind_rlucancel. X : Type a, b : X wlx0 : a = b
(fun (q : a = b) (s : wlx0 @ 1 = 1 @ q) =>
(fun (q0 : a = b) (s0 : q @ 1 = 1 @ q0) =>
let EH_x_y :=
((rlucancel_inv (rlucancel 1 [-] rlucancel 1 ))^ @
s) @ rlucancel_inv (rlucancel 1 [-] rlucancel 1 )
in
let EH_x_z :=
((rlucancel_inv (rlucancel 1 [-] rlucancel 1 ))^ @
s0) @ rlucancel_inv (rlucancel 1 [-] rlucancel 1 )
in
let EH_x_yz :=
((rlucancel_inv
(rlucancel 1
[-] (whiskerR 1 1 )^ @
(rlucancel 1 [-] rlucancel 1 )))^ @
((whiskerL wlx0 1 )^ @
((s [I] s0) @ whiskerR 1 q0))) @
rlucancel_inv
((whiskerR 1 1 )^ @ (rlucancel 1 [-] rlucancel 1 )
[-] rlucancel 1 ) in
EH_x_yz @ (concat_pp_p 1 1 q0 @ whiskerL 1 EH_x_z^) =
(concat_p_pp wlx0 1 1 @ whiskerR EH_x_y 1 ) @
concat_pp_p 1 q 1 ) q (rlucancel 1 )) wlx0
(rlucancel 1 )
destruct wlx0.X : Type a : X wlx0 : a = a
let EH_x_y :=
((rlucancel_inv (rlucancel 1 [-] rlucancel 1 ))^ @
rlucancel 1 ) @
rlucancel_inv (rlucancel 1 [-] rlucancel 1 ) in
let EH_x_z :=
((rlucancel_inv (rlucancel 1 [-] rlucancel 1 ))^ @
rlucancel 1 ) @
rlucancel_inv (rlucancel 1 [-] rlucancel 1 ) in
let EH_x_yz :=
((rlucancel_inv
(rlucancel 1
[-] (whiskerR 1 1 )^ @
(rlucancel 1 [-] rlucancel 1 )))^ @
((whiskerL 1 1 )^ @
((rlucancel 1 [I] rlucancel 1 ) @ whiskerR 1 1 ))) @
rlucancel_inv
((whiskerR 1 1 )^ @ (rlucancel 1 [-] rlucancel 1 )
[-] rlucancel 1 ) in
EH_x_yz @ (concat_pp_p 1 1 1 @ whiskerL 1 EH_x_z^) =
(concat_p_pp 1 1 1 @ whiskerR EH_x_y 1 ) @
concat_pp_p 1 1 1
clear wlx0.X : Type a : X
let EH_x_y :=
((rlucancel_inv (rlucancel 1 [-] rlucancel 1 ))^ @
rlucancel 1 ) @
rlucancel_inv (rlucancel 1 [-] rlucancel 1 ) in
let EH_x_z :=
((rlucancel_inv (rlucancel 1 [-] rlucancel 1 ))^ @
rlucancel 1 ) @
rlucancel_inv (rlucancel 1 [-] rlucancel 1 ) in
let EH_x_yz :=
((rlucancel_inv
(rlucancel 1
[-] (whiskerR 1 1 )^ @
(rlucancel 1 [-] rlucancel 1 )))^ @
((whiskerL 1 1 )^ @
((rlucancel 1 [I] rlucancel 1 ) @ whiskerR 1 1 ))) @
rlucancel_inv
((whiskerR 1 1 )^ @ (rlucancel 1 [-] rlucancel 1 )
[-] rlucancel 1 ) in
EH_x_yz @ (concat_pp_p 1 1 1 @ whiskerL 1 EH_x_z^) =
(concat_p_pp 1 1 1 @ whiskerR EH_x_y 1 ) @
concat_pp_p 1 1 1
reflexivity .
Defined .
End eh_p_pp .
Theorem eh_p_pp {X } {a : X} (p q r : idpath a = idpath a) :
eh p (q @ r) @ (concat_pp_p _ _ _ @ whiskerL _ (eh p r)^) =
concat_p_pp _ _ _ @ whiskerR (eh p q) _ @ concat_pp_p _ _ _.X : Type a : X p, q, r : 1 = 1
eh p (q @ r) @
(concat_pp_p q r p @ whiskerL q (eh p r)^) =
(concat_p_pp p q r @ whiskerR (eh p q) r) @
concat_pp_p q p r
Proof .X : Type a : X p, q, r : 1 = 1
eh p (q @ r) @
(concat_pp_p q r p @ whiskerL q (eh p r)^) =
(concat_p_pp p q r @ whiskerR (eh p q) r) @
concat_pp_p q p r
napply eh_p_pp_gen. X : Type a : X p, q, r : 1 = 1
urnat q [-] urnat r = whiskerR ?Goal 1 @ urnat (q @ r)
- X : Type a : X p, q, r : 1 = 1
urnat q [-] urnat r = whiskerR ?Goal 1 @ urnat (q @ r)
exact (urnat_pp q r).
- X : Type a : X p, q, r : 1 = 1
urnat q [-] urnat r = whiskerR ?Goal 1 @ urnat (q @ r)
exact (urnat_pp q r).
- X : Type a : X p, q, r : 1 = 1
(wlrnat p q [I] wlrnat p r) @
whiskerR (whiskerR_pp 1 q r)^ (whiskerL 1 p) =
whiskerL (whiskerL 1 p) (whiskerR_pp 1 q r)^ @
wlrnat p (q @ r)
exact (wlrnat_p_pp p q r).
Defined .
(* Coherence #1: We now prove that "eh (p @ q) r" suitably relates to "eh p r" and "eh q r". *)
Section eh_pp_p .
Context {X : Type }.
(* 0-paths *)
Context {a b c d e f : X}.
(* 1-paths *)
Context {wlx0 x0 : a = b}.
Context {wlx1 x1 : d = e}.
Context {wly0 y0 : b = c}.
Context {wly1 y1 : e = f}.
Context {wrz0 z0 : c = f}.
Context {wrz1 z1 : b = e}.
Context {wrz2 z2 : a = d}.
Context {wlxy0 : a = c}.
Context {wlxy1 : d = f}.
(* 2-paths *)
Context {ulnat_x0 : wlx0 @ 1 = 1 @ x0}.
Context {ulnat_x1 : wlx1 @ 1 = 1 @ x1}.
Context {ulnat_y0 : wly0 @ 1 = 1 @ y0}.
Context {ulnat_y1 : wly1 @ 1 = 1 @ y1}.
Context {urnat_z0 : wrz0 @ 1 = 1 @ z0}.
Context {urnat_z1 : wrz1 @ 1 = 1 @ z1}.
Context {urnat_z2 : wrz2 @ 1 = 1 @ z2}.
Context {ulnat_xy0 : wlxy0 @ 1 = 1 @ (x0 @ y0)}.
Context {ulnat_xy1 : wlxy1 @ 1 = 1 @ (x1 @ y1)}.
Context {wlrnat_x_z : wlx0 @ wrz1 = wrz2 @ wlx1}.
Context {wlrnat_y_z : wly0 @ wrz0 = wrz1 @ wly1}.
Context {wlrnat_xy_z : wlxy0 @ wrz0 = wrz2 @ wlxy1}.
Context {wlpp_xy0 : wlx0 @ wly0 = wlxy0}.
Context {wlpp_xy1 : wlx1 @ wly1 = wlxy1}.
(* 3-paths *)
Hypothesis H_ulnat_xy0 :
(ulnat_x0 [-] ulnat_y0) = whiskerR wlpp_xy0 _ @ ulnat_xy0.
Hypothesis H_ulnat_xy1 :
(ulnat_x1 [-] ulnat_y1) = whiskerR wlpp_xy1 _ @ ulnat_xy1.
Hypothesis H_wlrnat_xy_z :
(wlrnat_x_z [-] wlrnat_y_z) @ whiskerL _ wlpp_xy1 =
whiskerR wlpp_xy0 _ @ wlrnat_xy_z.
(* the coherence *)
Definition eh_pp_p_gen :
let EH_x_z := (rlucancel_inv (ulnat_x0 [-] urnat_z1))^ @
wlrnat_x_z @ rlucancel_inv (urnat_z2 [-] ulnat_x1) in
let EH_y_z := (rlucancel_inv (ulnat_y0 [-] urnat_z0))^ @
wlrnat_y_z @ rlucancel_inv (urnat_z1 [-] ulnat_y1) in
let EH_xy_z := (rlucancel_inv (ulnat_xy0 [-] urnat_z0))^ @
wlrnat_xy_z @ rlucancel_inv (urnat_z2 [-] ulnat_xy1) in
EH_xy_z @ (concat_p_pp _ _ _ @ whiskerR EH_x_z^ _) =
concat_pp_p _ _ _ @ whiskerL _ EH_y_z @ concat_p_pp _ _ _.X : Type a, b, c, d, e, f : X wlx0, x0 : a = b wlx1, x1 : d = e wly0, y0 : b = c wly1, y1 : e = f wrz0, z0 : c = f wrz1, z1 : b = e wrz2, z2 : a = d wlxy0 : a = c wlxy1 : d = f ulnat_x0 : wlx0 @ 1 = 1 @ x0 ulnat_x1 : wlx1 @ 1 = 1 @ x1 ulnat_y0 : wly0 @ 1 = 1 @ y0 ulnat_y1 : wly1 @ 1 = 1 @ y1 urnat_z0 : wrz0 @ 1 = 1 @ z0 urnat_z1 : wrz1 @ 1 = 1 @ z1 urnat_z2 : wrz2 @ 1 = 1 @ z2 ulnat_xy0 : wlxy0 @ 1 = 1 @ (x0 @ y0) ulnat_xy1 : wlxy1 @ 1 = 1 @ (x1 @ y1) wlrnat_x_z : wlx0 @ wrz1 = wrz2 @ wlx1 wlrnat_y_z : wly0 @ wrz0 = wrz1 @ wly1 wlrnat_xy_z : wlxy0 @ wrz0 = wrz2 @ wlxy1 wlpp_xy0 : wlx0 @ wly0 = wlxy0 wlpp_xy1 : wlx1 @ wly1 = wlxy1 H_ulnat_xy0 : ulnat_x0 [-] ulnat_y0 =
whiskerR wlpp_xy0 1 @ ulnat_xy0 H_ulnat_xy1 : ulnat_x1 [-] ulnat_y1 =
whiskerR wlpp_xy1 1 @ ulnat_xy1 H_wlrnat_xy_z : (wlrnat_x_z [-] wlrnat_y_z) @
whiskerL wrz2 wlpp_xy1 =
whiskerR wlpp_xy0 wrz0 @ wlrnat_xy_z
let EH_x_z :=
((rlucancel_inv (ulnat_x0 [-] urnat_z1))^ @
wlrnat_x_z) @ rlucancel_inv (urnat_z2 [-] ulnat_x1)
in
let EH_y_z :=
((rlucancel_inv (ulnat_y0 [-] urnat_z0))^ @
wlrnat_y_z) @ rlucancel_inv (urnat_z1 [-] ulnat_y1)
in
let EH_xy_z :=
((rlucancel_inv (ulnat_xy0 [-] urnat_z0))^ @
wlrnat_xy_z) @
rlucancel_inv (urnat_z2 [-] ulnat_xy1) in
EH_xy_z @ (concat_p_pp z2 x1 y1 @ whiskerR EH_x_z^ y1) =
(concat_pp_p x0 y0 z0 @ whiskerL x0 EH_y_z) @
concat_p_pp x0 z1 y1
Proof .X : Type a, b, c, d, e, f : X wlx0, x0 : a = b wlx1, x1 : d = e wly0, y0 : b = c wly1, y1 : e = f wrz0, z0 : c = f wrz1, z1 : b = e wrz2, z2 : a = d wlxy0 : a = c wlxy1 : d = f ulnat_x0 : wlx0 @ 1 = 1 @ x0 ulnat_x1 : wlx1 @ 1 = 1 @ x1 ulnat_y0 : wly0 @ 1 = 1 @ y0 ulnat_y1 : wly1 @ 1 = 1 @ y1 urnat_z0 : wrz0 @ 1 = 1 @ z0 urnat_z1 : wrz1 @ 1 = 1 @ z1 urnat_z2 : wrz2 @ 1 = 1 @ z2 ulnat_xy0 : wlxy0 @ 1 = 1 @ (x0 @ y0) ulnat_xy1 : wlxy1 @ 1 = 1 @ (x1 @ y1) wlrnat_x_z : wlx0 @ wrz1 = wrz2 @ wlx1 wlrnat_y_z : wly0 @ wrz0 = wrz1 @ wly1 wlrnat_xy_z : wlxy0 @ wrz0 = wrz2 @ wlxy1 wlpp_xy0 : wlx0 @ wly0 = wlxy0 wlpp_xy1 : wlx1 @ wly1 = wlxy1 H_ulnat_xy0 : ulnat_x0 [-] ulnat_y0 =
whiskerR wlpp_xy0 1 @ ulnat_xy0 H_ulnat_xy1 : ulnat_x1 [-] ulnat_y1 =
whiskerR wlpp_xy1 1 @ ulnat_xy1 H_wlrnat_xy_z : (wlrnat_x_z [-] wlrnat_y_z) @
whiskerL wrz2 wlpp_xy1 =
whiskerR wlpp_xy0 wrz0 @ wlrnat_xy_z
let EH_x_z :=
((rlucancel_inv (ulnat_x0 [-] urnat_z1))^ @
wlrnat_x_z) @ rlucancel_inv (urnat_z2 [-] ulnat_x1)
in
let EH_y_z :=
((rlucancel_inv (ulnat_y0 [-] urnat_z0))^ @
wlrnat_y_z) @ rlucancel_inv (urnat_z1 [-] ulnat_y1)
in
let EH_xy_z :=
((rlucancel_inv (ulnat_xy0 [-] urnat_z0))^ @
wlrnat_xy_z) @
rlucancel_inv (urnat_z2 [-] ulnat_xy1) in
EH_xy_z @ (concat_p_pp z2 x1 y1 @ whiskerR EH_x_z^ y1) =
(concat_pp_p x0 y0 z0 @ whiskerL x0 EH_y_z) @
concat_p_pp x0 z1 y1
apply moveR_Vp in H_ulnat_xy0, H_ulnat_xy1, H_wlrnat_xy_z.X : Type a, b, c, d, e, f : X wlx0, x0 : a = b wlx1, x1 : d = e wly0, y0 : b = c wly1, y1 : e = f wrz0, z0 : c = f wrz1, z1 : b = e wrz2, z2 : a = d wlxy0 : a = c wlxy1 : d = f ulnat_x0 : wlx0 @ 1 = 1 @ x0 ulnat_x1 : wlx1 @ 1 = 1 @ x1 ulnat_y0 : wly0 @ 1 = 1 @ y0 ulnat_y1 : wly1 @ 1 = 1 @ y1 urnat_z0 : wrz0 @ 1 = 1 @ z0 urnat_z1 : wrz1 @ 1 = 1 @ z1 urnat_z2 : wrz2 @ 1 = 1 @ z2 ulnat_xy0 : wlxy0 @ 1 = 1 @ (x0 @ y0) ulnat_xy1 : wlxy1 @ 1 = 1 @ (x1 @ y1) wlrnat_x_z : wlx0 @ wrz1 = wrz2 @ wlx1 wlrnat_y_z : wly0 @ wrz0 = wrz1 @ wly1 wlrnat_xy_z : wlxy0 @ wrz0 = wrz2 @ wlxy1 wlpp_xy0 : wlx0 @ wly0 = wlxy0 wlpp_xy1 : wlx1 @ wly1 = wlxy1 H_ulnat_xy0 : (whiskerR wlpp_xy0 1 )^ @
(ulnat_x0 [-] ulnat_y0) = ulnat_xy0 H_ulnat_xy1 : (whiskerR wlpp_xy1 1 )^ @
(ulnat_x1 [-] ulnat_y1) = ulnat_xy1 H_wlrnat_xy_z : (whiskerR wlpp_xy0 wrz0)^ @
((wlrnat_x_z [-] wlrnat_y_z) @
whiskerL wrz2 wlpp_xy1) = wlrnat_xy_z
let EH_x_z :=
((rlucancel_inv (ulnat_x0 [-] urnat_z1))^ @
wlrnat_x_z) @ rlucancel_inv (urnat_z2 [-] ulnat_x1)
in
let EH_y_z :=
((rlucancel_inv (ulnat_y0 [-] urnat_z0))^ @
wlrnat_y_z) @ rlucancel_inv (urnat_z1 [-] ulnat_y1)
in
let EH_xy_z :=
((rlucancel_inv (ulnat_xy0 [-] urnat_z0))^ @
wlrnat_xy_z) @
rlucancel_inv (urnat_z2 [-] ulnat_xy1) in
EH_xy_z @ (concat_p_pp z2 x1 y1 @ whiskerR EH_x_z^ y1) =
(concat_pp_p x0 y0 z0 @ whiskerL x0 EH_y_z) @
concat_p_pp x0 z1 y1
destruct H_ulnat_xy0, H_ulnat_xy1, H_wlrnat_xy_z.X : Type a, b, c, d, e, f : X wlx0, x0 : a = b wlx1, x1 : d = e wly0, y0 : b = c wly1, y1 : e = f wrz0, z0 : c = f wrz1, z1 : b = e wrz2, z2 : a = d wlxy0 : a = c wlxy1 : d = f ulnat_x0 : wlx0 @ 1 = 1 @ x0 ulnat_x1 : wlx1 @ 1 = 1 @ x1 ulnat_y0 : wly0 @ 1 = 1 @ y0 ulnat_y1 : wly1 @ 1 = 1 @ y1 urnat_z0 : wrz0 @ 1 = 1 @ z0 urnat_z1 : wrz1 @ 1 = 1 @ z1 urnat_z2 : wrz2 @ 1 = 1 @ z2 wlrnat_x_z : wlx0 @ wrz1 = wrz2 @ wlx1 wlrnat_y_z : wly0 @ wrz0 = wrz1 @ wly1 wlpp_xy0 : wlx0 @ wly0 = wlxy0 wlpp_xy1 : wlx1 @ wly1 = wlxy1 H_ulnat_xy0 : (whiskerR wlpp_xy0 1 )^ @
(ulnat_x0 [-] ulnat_y0) =
(whiskerR wlpp_xy0 1 )^ @
(ulnat_x0 [-] ulnat_y0) H_ulnat_xy1 : (whiskerR wlpp_xy1 1 )^ @
(ulnat_x1 [-] ulnat_y1) =
(whiskerR wlpp_xy1 1 )^ @
(ulnat_x1 [-] ulnat_y1) H_wlrnat_xy_z : (whiskerR wlpp_xy0 wrz0)^ @
((wlrnat_x_z [-] wlrnat_y_z) @
whiskerL wrz2 wlpp_xy1) =
(whiskerR wlpp_xy0 wrz0)^ @
((wlrnat_x_z [-] wlrnat_y_z) @
whiskerL wrz2 wlpp_xy1)
let EH_x_z :=
((rlucancel_inv (ulnat_x0 [-] urnat_z1))^ @
wlrnat_x_z) @ rlucancel_inv (urnat_z2 [-] ulnat_x1)
in
let EH_y_z :=
((rlucancel_inv (ulnat_y0 [-] urnat_z0))^ @
wlrnat_y_z) @ rlucancel_inv (urnat_z1 [-] ulnat_y1)
in
let EH_xy_z :=
((rlucancel_inv
((whiskerR wlpp_xy0 1 )^ @
(ulnat_x0 [-] ulnat_y0) [-] urnat_z0))^ @
((whiskerR wlpp_xy0 wrz0)^ @
((wlrnat_x_z [-] wlrnat_y_z) @
whiskerL wrz2 wlpp_xy1))) @
rlucancel_inv
(urnat_z2
[-] (whiskerR wlpp_xy1 1 )^ @
(ulnat_x1 [-] ulnat_y1)) in
EH_xy_z @ (concat_p_pp z2 x1 y1 @ whiskerR EH_x_z^ y1) =
(concat_pp_p x0 y0 z0 @ whiskerL x0 EH_y_z) @
concat_p_pp x0 z1 y1
clear H_ulnat_xy0 H_ulnat_xy1 H_wlrnat_xy_z.X : Type a, b, c, d, e, f : X wlx0, x0 : a = b wlx1, x1 : d = e wly0, y0 : b = c wly1, y1 : e = f wrz0, z0 : c = f wrz1, z1 : b = e wrz2, z2 : a = d wlxy0 : a = c wlxy1 : d = f ulnat_x0 : wlx0 @ 1 = 1 @ x0 ulnat_x1 : wlx1 @ 1 = 1 @ x1 ulnat_y0 : wly0 @ 1 = 1 @ y0 ulnat_y1 : wly1 @ 1 = 1 @ y1 urnat_z0 : wrz0 @ 1 = 1 @ z0 urnat_z1 : wrz1 @ 1 = 1 @ z1 urnat_z2 : wrz2 @ 1 = 1 @ z2 wlrnat_x_z : wlx0 @ wrz1 = wrz2 @ wlx1 wlrnat_y_z : wly0 @ wrz0 = wrz1 @ wly1 wlpp_xy0 : wlx0 @ wly0 = wlxy0 wlpp_xy1 : wlx1 @ wly1 = wlxy1
let EH_x_z :=
((rlucancel_inv (ulnat_x0 [-] urnat_z1))^ @
wlrnat_x_z) @ rlucancel_inv (urnat_z2 [-] ulnat_x1)
in
let EH_y_z :=
((rlucancel_inv (ulnat_y0 [-] urnat_z0))^ @
wlrnat_y_z) @ rlucancel_inv (urnat_z1 [-] ulnat_y1)
in
let EH_xy_z :=
((rlucancel_inv
((whiskerR wlpp_xy0 1 )^ @
(ulnat_x0 [-] ulnat_y0) [-] urnat_z0))^ @
((whiskerR wlpp_xy0 wrz0)^ @
((wlrnat_x_z [-] wlrnat_y_z) @
whiskerL wrz2 wlpp_xy1))) @
rlucancel_inv
(urnat_z2
[-] (whiskerR wlpp_xy1 1 )^ @
(ulnat_x1 [-] ulnat_y1)) in
EH_xy_z @ (concat_p_pp z2 x1 y1 @ whiskerR EH_x_z^ y1) =
(concat_pp_p x0 y0 z0 @ whiskerL x0 EH_y_z) @
concat_p_pp x0 z1 y1
destruct wlpp_xy0, wlpp_xy1.X : Type a, b, c, d, e, f : X wlx0, x0 : a = b wlx1, x1 : d = e wly0, y0 : b = c wly1, y1 : e = f wrz0, z0 : c = f wrz1, z1 : b = e wrz2, z2 : a = d ulnat_x0 : wlx0 @ 1 = 1 @ x0 ulnat_x1 : wlx1 @ 1 = 1 @ x1 ulnat_y0 : wly0 @ 1 = 1 @ y0 ulnat_y1 : wly1 @ 1 = 1 @ y1 urnat_z0 : wrz0 @ 1 = 1 @ z0 urnat_z1 : wrz1 @ 1 = 1 @ z1 urnat_z2 : wrz2 @ 1 = 1 @ z2 wlrnat_x_z : wlx0 @ wrz1 = wrz2 @ wlx1 wlrnat_y_z : wly0 @ wrz0 = wrz1 @ wly1 wlpp_xy0 : wlx0 @ wly0 = wlx0 @ wly0 wlpp_xy1 : wlx1 @ wly1 = wlx1 @ wly1
let EH_x_z :=
((rlucancel_inv (ulnat_x0 [-] urnat_z1))^ @
wlrnat_x_z) @ rlucancel_inv (urnat_z2 [-] ulnat_x1)
in
let EH_y_z :=
((rlucancel_inv (ulnat_y0 [-] urnat_z0))^ @
wlrnat_y_z) @ rlucancel_inv (urnat_z1 [-] ulnat_y1)
in
let EH_xy_z :=
((rlucancel_inv
((whiskerR 1 1 )^ @ (ulnat_x0 [-] ulnat_y0)
[-] urnat_z0))^ @
((whiskerR 1 wrz0)^ @
((wlrnat_x_z [-] wlrnat_y_z) @ whiskerL wrz2 1 ))) @
rlucancel_inv
(urnat_z2
[-] (whiskerR 1 1 )^ @ (ulnat_x1 [-] ulnat_y1)) in
EH_xy_z @ (concat_p_pp z2 x1 y1 @ whiskerR EH_x_z^ y1) =
(concat_pp_p x0 y0 z0 @ whiskerL x0 EH_y_z) @
concat_p_pp x0 z1 y1
clear wlpp_xy0 wlpp_xy1.X : Type a, b, c, d, e, f : X wlx0, x0 : a = b wlx1, x1 : d = e wly0, y0 : b = c wly1, y1 : e = f wrz0, z0 : c = f wrz1, z1 : b = e wrz2, z2 : a = d ulnat_x0 : wlx0 @ 1 = 1 @ x0 ulnat_x1 : wlx1 @ 1 = 1 @ x1 ulnat_y0 : wly0 @ 1 = 1 @ y0 ulnat_y1 : wly1 @ 1 = 1 @ y1 urnat_z0 : wrz0 @ 1 = 1 @ z0 urnat_z1 : wrz1 @ 1 = 1 @ z1 urnat_z2 : wrz2 @ 1 = 1 @ z2 wlrnat_x_z : wlx0 @ wrz1 = wrz2 @ wlx1 wlrnat_y_z : wly0 @ wrz0 = wrz1 @ wly1
let EH_x_z :=
((rlucancel_inv (ulnat_x0 [-] urnat_z1))^ @
wlrnat_x_z) @ rlucancel_inv (urnat_z2 [-] ulnat_x1)
in
let EH_y_z :=
((rlucancel_inv (ulnat_y0 [-] urnat_z0))^ @
wlrnat_y_z) @ rlucancel_inv (urnat_z1 [-] ulnat_y1)
in
let EH_xy_z :=
((rlucancel_inv
((whiskerR 1 1 )^ @ (ulnat_x0 [-] ulnat_y0)
[-] urnat_z0))^ @
((whiskerR 1 wrz0)^ @
((wlrnat_x_z [-] wlrnat_y_z) @ whiskerL wrz2 1 ))) @
rlucancel_inv
(urnat_z2
[-] (whiskerR 1 1 )^ @ (ulnat_x1 [-] ulnat_y1)) in
EH_xy_z @ (concat_p_pp z2 x1 y1 @ whiskerR EH_x_z^ y1) =
(concat_pp_p x0 y0 z0 @ whiskerL x0 EH_y_z) @
concat_p_pp x0 z1 y1
revert x0 ulnat_x0.X : Type a, b, c, d, e, f : X wlx0 : a = b wlx1, x1 : d = e wly0, y0 : b = c wly1, y1 : e = f wrz0, z0 : c = f wrz1, z1 : b = e wrz2, z2 : a = d ulnat_x1 : wlx1 @ 1 = 1 @ x1 ulnat_y0 : wly0 @ 1 = 1 @ y0 ulnat_y1 : wly1 @ 1 = 1 @ y1 urnat_z0 : wrz0 @ 1 = 1 @ z0 urnat_z1 : wrz1 @ 1 = 1 @ z1 urnat_z2 : wrz2 @ 1 = 1 @ z2 wlrnat_x_z : wlx0 @ wrz1 = wrz2 @ wlx1 wlrnat_y_z : wly0 @ wrz0 = wrz1 @ wly1
forall (x0 : a = b) (ulnat_x0 : wlx0 @ 1 = 1 @ x0),
let EH_x_z :=
((rlucancel_inv (ulnat_x0 [-] urnat_z1))^ @
wlrnat_x_z) @ rlucancel_inv (urnat_z2 [-] ulnat_x1)
in
let EH_y_z :=
((rlucancel_inv (ulnat_y0 [-] urnat_z0))^ @
wlrnat_y_z) @ rlucancel_inv (urnat_z1 [-] ulnat_y1)
in
let EH_xy_z :=
((rlucancel_inv
((whiskerR 1 1 )^ @ (ulnat_x0 [-] ulnat_y0)
[-] urnat_z0))^ @
((whiskerR 1 wrz0)^ @
((wlrnat_x_z [-] wlrnat_y_z) @ whiskerL wrz2 1 ))) @
rlucancel_inv
(urnat_z2
[-] (whiskerR 1 1 )^ @ (ulnat_x1 [-] ulnat_y1)) in
EH_xy_z @ (concat_p_pp z2 x1 y1 @ whiskerR EH_x_z^ y1) =
(concat_pp_p x0 y0 z0 @ whiskerL x0 EH_y_z) @
concat_p_pp x0 z1 y1
snapply equiv_path_ind_rlucancel. X : Type a, b, c, d, e, f : X wlx0 : a = b wlx1, x1 : d = e wly0, y0 : b = c wly1, y1 : e = f wrz0, z0 : c = f wrz1, z1 : b = e wrz2, z2 : a = d ulnat_x1 : wlx1 @ 1 = 1 @ x1 ulnat_y0 : wly0 @ 1 = 1 @ y0 ulnat_y1 : wly1 @ 1 = 1 @ y1 urnat_z0 : wrz0 @ 1 = 1 @ z0 urnat_z1 : wrz1 @ 1 = 1 @ z1 urnat_z2 : wrz2 @ 1 = 1 @ z2 wlrnat_x_z : wlx0 @ wrz1 = wrz2 @ wlx1 wlrnat_y_z : wly0 @ wrz0 = wrz1 @ wly1
(fun (q : a = b) (s : wlx0 @ 1 = 1 @ q) =>
let EH_x_z :=
((rlucancel_inv (s [-] urnat_z1))^ @ wlrnat_x_z) @
rlucancel_inv (urnat_z2 [-] ulnat_x1) in
let EH_y_z :=
((rlucancel_inv (ulnat_y0 [-] urnat_z0))^ @
wlrnat_y_z) @
rlucancel_inv (urnat_z1 [-] ulnat_y1) in
let EH_xy_z :=
((rlucancel_inv
((whiskerR 1 1 )^ @ (s [-] ulnat_y0)
[-] urnat_z0))^ @
((whiskerR 1 wrz0)^ @
((wlrnat_x_z [-] wlrnat_y_z) @ whiskerL wrz2 1 ))) @
rlucancel_inv
(urnat_z2
[-] (whiskerR 1 1 )^ @ (ulnat_x1 [-] ulnat_y1))
in
EH_xy_z @
(concat_p_pp z2 x1 y1 @ whiskerR EH_x_z^ y1) =
(concat_pp_p q y0 z0 @ whiskerL q EH_y_z) @
concat_p_pp q z1 y1) wlx0 (rlucancel 1 )
revert x1 ulnat_x1.X : Type a, b, c, d, e, f : X wlx0 : a = b wlx1 : d = e wly0, y0 : b = c wly1, y1 : e = f wrz0, z0 : c = f wrz1, z1 : b = e wrz2, z2 : a = d ulnat_y0 : wly0 @ 1 = 1 @ y0 ulnat_y1 : wly1 @ 1 = 1 @ y1 urnat_z0 : wrz0 @ 1 = 1 @ z0 urnat_z1 : wrz1 @ 1 = 1 @ z1 urnat_z2 : wrz2 @ 1 = 1 @ z2 wlrnat_x_z : wlx0 @ wrz1 = wrz2 @ wlx1 wlrnat_y_z : wly0 @ wrz0 = wrz1 @ wly1
forall (x1 : d = e) (ulnat_x1 : wlx1 @ 1 = 1 @ x1),
(fun (q : a = b) (s : wlx0 @ 1 = 1 @ q) =>
let EH_x_z :=
((rlucancel_inv (s [-] urnat_z1))^ @ wlrnat_x_z) @
rlucancel_inv (urnat_z2 [-] ulnat_x1) in
let EH_y_z :=
((rlucancel_inv (ulnat_y0 [-] urnat_z0))^ @
wlrnat_y_z) @
rlucancel_inv (urnat_z1 [-] ulnat_y1) in
let EH_xy_z :=
((rlucancel_inv
((whiskerR 1 1 )^ @ (s [-] ulnat_y0)
[-] urnat_z0))^ @
((whiskerR 1 wrz0)^ @
((wlrnat_x_z [-] wlrnat_y_z) @ whiskerL wrz2 1 ))) @
rlucancel_inv
(urnat_z2
[-] (whiskerR 1 1 )^ @ (ulnat_x1 [-] ulnat_y1))
in
EH_xy_z @
(concat_p_pp z2 x1 y1 @ whiskerR EH_x_z^ y1) =
(concat_pp_p q y0 z0 @ whiskerL q EH_y_z) @
concat_p_pp q z1 y1) wlx0 (rlucancel 1 )
snapply equiv_path_ind_rlucancel. X : Type a, b, c, d, e, f : X wlx0 : a = b wlx1 : d = e wly0, y0 : b = c wly1, y1 : e = f wrz0, z0 : c = f wrz1, z1 : b = e wrz2, z2 : a = d ulnat_y0 : wly0 @ 1 = 1 @ y0 ulnat_y1 : wly1 @ 1 = 1 @ y1 urnat_z0 : wrz0 @ 1 = 1 @ z0 urnat_z1 : wrz1 @ 1 = 1 @ z1 urnat_z2 : wrz2 @ 1 = 1 @ z2 wlrnat_x_z : wlx0 @ wrz1 = wrz2 @ wlx1 wlrnat_y_z : wly0 @ wrz0 = wrz1 @ wly1
(fun (q : d = e) (s : wlx1 @ 1 = 1 @ q) =>
(fun (q0 : a = b) (s0 : wlx0 @ 1 = 1 @ q0) =>
let EH_x_z :=
((rlucancel_inv (s0 [-] urnat_z1))^ @ wlrnat_x_z) @
rlucancel_inv (urnat_z2 [-] s) in
let EH_y_z :=
((rlucancel_inv (ulnat_y0 [-] urnat_z0))^ @
wlrnat_y_z) @
rlucancel_inv (urnat_z1 [-] ulnat_y1) in
let EH_xy_z :=
((rlucancel_inv
((whiskerR 1 1 )^ @ (s0 [-] ulnat_y0)
[-] urnat_z0))^ @
((whiskerR 1 wrz0)^ @
((wlrnat_x_z [-] wlrnat_y_z) @ whiskerL wrz2 1 ))) @
rlucancel_inv
(urnat_z2 [-] (whiskerR 1 1 )^ @ (s [-] ulnat_y1))
in
EH_xy_z @
(concat_p_pp z2 q y1 @ whiskerR EH_x_z^ y1) =
(concat_pp_p q0 y0 z0 @ whiskerL q0 EH_y_z) @
concat_p_pp q0 z1 y1) wlx0 (rlucancel 1 )) wlx1
(rlucancel 1 )
revert y0 ulnat_y0.X : Type a, b, c, d, e, f : X wlx0 : a = b wlx1 : d = e wly0 : b = c wly1, y1 : e = f wrz0, z0 : c = f wrz1, z1 : b = e wrz2, z2 : a = d ulnat_y1 : wly1 @ 1 = 1 @ y1 urnat_z0 : wrz0 @ 1 = 1 @ z0 urnat_z1 : wrz1 @ 1 = 1 @ z1 urnat_z2 : wrz2 @ 1 = 1 @ z2 wlrnat_x_z : wlx0 @ wrz1 = wrz2 @ wlx1 wlrnat_y_z : wly0 @ wrz0 = wrz1 @ wly1
forall (y0 : b = c) (ulnat_y0 : wly0 @ 1 = 1 @ y0),
(fun (q : d = e) (s : wlx1 @ 1 = 1 @ q) =>
(fun (q0 : a = b) (s0 : wlx0 @ 1 = 1 @ q0) =>
let EH_x_z :=
((rlucancel_inv (s0 [-] urnat_z1))^ @ wlrnat_x_z) @
rlucancel_inv (urnat_z2 [-] s) in
let EH_y_z :=
((rlucancel_inv (ulnat_y0 [-] urnat_z0))^ @
wlrnat_y_z) @
rlucancel_inv (urnat_z1 [-] ulnat_y1) in
let EH_xy_z :=
((rlucancel_inv
((whiskerR 1 1 )^ @ (s0 [-] ulnat_y0)
[-] urnat_z0))^ @
((whiskerR 1 wrz0)^ @
((wlrnat_x_z [-] wlrnat_y_z) @ whiskerL wrz2 1 ))) @
rlucancel_inv
(urnat_z2 [-] (whiskerR 1 1 )^ @ (s [-] ulnat_y1))
in
EH_xy_z @
(concat_p_pp z2 q y1 @ whiskerR EH_x_z^ y1) =
(concat_pp_p q0 y0 z0 @ whiskerL q0 EH_y_z) @
concat_p_pp q0 z1 y1) wlx0 (rlucancel 1 )) wlx1
(rlucancel 1 )
snapply equiv_path_ind_rlucancel. X : Type a, b, c, d, e, f : X wlx0 : a = b wlx1 : d = e wly0 : b = c wly1, y1 : e = f wrz0, z0 : c = f wrz1, z1 : b = e wrz2, z2 : a = d ulnat_y1 : wly1 @ 1 = 1 @ y1 urnat_z0 : wrz0 @ 1 = 1 @ z0 urnat_z1 : wrz1 @ 1 = 1 @ z1 urnat_z2 : wrz2 @ 1 = 1 @ z2 wlrnat_x_z : wlx0 @ wrz1 = wrz2 @ wlx1 wlrnat_y_z : wly0 @ wrz0 = wrz1 @ wly1
(fun (q : b = c) (s : wly0 @ 1 = 1 @ q) =>
(fun (q0 : d = e) (s0 : wlx1 @ 1 = 1 @ q0) =>
(fun (q1 : a = b) (s1 : wlx0 @ 1 = 1 @ q1) =>
let EH_x_z :=
((rlucancel_inv (s1 [-] urnat_z1))^ @ wlrnat_x_z) @
rlucancel_inv (urnat_z2 [-] s0) in
let EH_y_z :=
((rlucancel_inv (s [-] urnat_z0))^ @ wlrnat_y_z) @
rlucancel_inv (urnat_z1 [-] ulnat_y1) in
let EH_xy_z :=
((rlucancel_inv
((whiskerR 1 1 )^ @ (s1 [-] s) [-] urnat_z0))^ @
((whiskerR 1 wrz0)^ @
((wlrnat_x_z [-] wlrnat_y_z) @ whiskerL wrz2 1 ))) @
rlucancel_inv
(urnat_z2
[-] (whiskerR 1 1 )^ @ (s0 [-] ulnat_y1)) in
EH_xy_z @
(concat_p_pp z2 q0 y1 @ whiskerR EH_x_z^ y1) =
(concat_pp_p q1 q z0 @ whiskerL q1 EH_y_z) @
concat_p_pp q1 z1 y1) wlx0 (rlucancel 1 )) wlx1
(rlucancel 1 )) wly0 (rlucancel 1 )
revert y1 ulnat_y1.X : Type a, b, c, d, e, f : X wlx0 : a = b wlx1 : d = e wly0 : b = c wly1 : e = f wrz0, z0 : c = f wrz1, z1 : b = e wrz2, z2 : a = d urnat_z0 : wrz0 @ 1 = 1 @ z0 urnat_z1 : wrz1 @ 1 = 1 @ z1 urnat_z2 : wrz2 @ 1 = 1 @ z2 wlrnat_x_z : wlx0 @ wrz1 = wrz2 @ wlx1 wlrnat_y_z : wly0 @ wrz0 = wrz1 @ wly1
forall (y1 : e = f) (ulnat_y1 : wly1 @ 1 = 1 @ y1),
(fun (q : b = c) (s : wly0 @ 1 = 1 @ q) =>
(fun (q0 : d = e) (s0 : wlx1 @ 1 = 1 @ q0) =>
(fun (q1 : a = b) (s1 : wlx0 @ 1 = 1 @ q1) =>
let EH_x_z :=
((rlucancel_inv (s1 [-] urnat_z1))^ @ wlrnat_x_z) @
rlucancel_inv (urnat_z2 [-] s0) in
let EH_y_z :=
((rlucancel_inv (s [-] urnat_z0))^ @ wlrnat_y_z) @
rlucancel_inv (urnat_z1 [-] ulnat_y1) in
let EH_xy_z :=
((rlucancel_inv
((whiskerR 1 1 )^ @ (s1 [-] s) [-] urnat_z0))^ @
((whiskerR 1 wrz0)^ @
((wlrnat_x_z [-] wlrnat_y_z) @ whiskerL wrz2 1 ))) @
rlucancel_inv
(urnat_z2
[-] (whiskerR 1 1 )^ @ (s0 [-] ulnat_y1)) in
EH_xy_z @
(concat_p_pp z2 q0 y1 @ whiskerR EH_x_z^ y1) =
(concat_pp_p q1 q z0 @ whiskerL q1 EH_y_z) @
concat_p_pp q1 z1 y1) wlx0 (rlucancel 1 )) wlx1
(rlucancel 1 )) wly0 (rlucancel 1 )
snapply equiv_path_ind_rlucancel. X : Type a, b, c, d, e, f : X wlx0 : a = b wlx1 : d = e wly0 : b = c wly1 : e = f wrz0, z0 : c = f wrz1, z1 : b = e wrz2, z2 : a = d urnat_z0 : wrz0 @ 1 = 1 @ z0 urnat_z1 : wrz1 @ 1 = 1 @ z1 urnat_z2 : wrz2 @ 1 = 1 @ z2 wlrnat_x_z : wlx0 @ wrz1 = wrz2 @ wlx1 wlrnat_y_z : wly0 @ wrz0 = wrz1 @ wly1
(fun (q : e = f) (s : wly1 @ 1 = 1 @ q) =>
(fun (q0 : b = c) (s0 : wly0 @ 1 = 1 @ q0) =>
(fun (q1 : d = e) (s1 : wlx1 @ 1 = 1 @ q1) =>
(fun (q2 : a = b) (s2 : wlx0 @ 1 = 1 @ q2) =>
let EH_x_z :=
((rlucancel_inv (s2 [-] urnat_z1))^ @ wlrnat_x_z) @
rlucancel_inv (urnat_z2 [-] s1) in
let EH_y_z :=
((rlucancel_inv (s0 [-] urnat_z0))^ @ wlrnat_y_z) @
rlucancel_inv (urnat_z1 [-] s) in
let EH_xy_z :=
((rlucancel_inv
((whiskerR 1 1 )^ @ (s2 [-] s0) [-] urnat_z0))^ @
((whiskerR 1 wrz0)^ @
((wlrnat_x_z [-] wlrnat_y_z) @ whiskerL wrz2 1 ))) @
rlucancel_inv
(urnat_z2 [-] (whiskerR 1 1 )^ @ (s1 [-] s)) in
EH_xy_z @
(concat_p_pp z2 q1 q @ whiskerR EH_x_z^ q) =
(concat_pp_p q2 q0 z0 @ whiskerL q2 EH_y_z) @
concat_p_pp q2 z1 q) wlx0 (rlucancel 1 )) wlx1
(rlucancel 1 )) wly0 (rlucancel 1 )) wly1
(rlucancel 1 )
revert z0 urnat_z0.X : Type a, b, c, d, e, f : X wlx0 : a = b wlx1 : d = e wly0 : b = c wly1 : e = f wrz0 : c = f wrz1, z1 : b = e wrz2, z2 : a = d urnat_z1 : wrz1 @ 1 = 1 @ z1 urnat_z2 : wrz2 @ 1 = 1 @ z2 wlrnat_x_z : wlx0 @ wrz1 = wrz2 @ wlx1 wlrnat_y_z : wly0 @ wrz0 = wrz1 @ wly1
forall (z0 : c = f) (urnat_z0 : wrz0 @ 1 = 1 @ z0),
(fun (q : e = f) (s : wly1 @ 1 = 1 @ q) =>
(fun (q0 : b = c) (s0 : wly0 @ 1 = 1 @ q0) =>
(fun (q1 : d = e) (s1 : wlx1 @ 1 = 1 @ q1) =>
(fun (q2 : a = b) (s2 : wlx0 @ 1 = 1 @ q2) =>
let EH_x_z :=
((rlucancel_inv (s2 [-] urnat_z1))^ @ wlrnat_x_z) @
rlucancel_inv (urnat_z2 [-] s1) in
let EH_y_z :=
((rlucancel_inv (s0 [-] urnat_z0))^ @ wlrnat_y_z) @
rlucancel_inv (urnat_z1 [-] s) in
let EH_xy_z :=
((rlucancel_inv
((whiskerR 1 1 )^ @ (s2 [-] s0) [-] urnat_z0))^ @
((whiskerR 1 wrz0)^ @
((wlrnat_x_z [-] wlrnat_y_z) @ whiskerL wrz2 1 ))) @
rlucancel_inv
(urnat_z2 [-] (whiskerR 1 1 )^ @ (s1 [-] s)) in
EH_xy_z @
(concat_p_pp z2 q1 q @ whiskerR EH_x_z^ q) =
(concat_pp_p q2 q0 z0 @ whiskerL q2 EH_y_z) @
concat_p_pp q2 z1 q) wlx0 (rlucancel 1 )) wlx1
(rlucancel 1 )) wly0 (rlucancel 1 )) wly1
(rlucancel 1 )
snapply equiv_path_ind_rlucancel. X : Type a, b, c, d, e, f : X wlx0 : a = b wlx1 : d = e wly0 : b = c wly1 : e = f wrz0 : c = f wrz1, z1 : b = e wrz2, z2 : a = d urnat_z1 : wrz1 @ 1 = 1 @ z1 urnat_z2 : wrz2 @ 1 = 1 @ z2 wlrnat_x_z : wlx0 @ wrz1 = wrz2 @ wlx1 wlrnat_y_z : wly0 @ wrz0 = wrz1 @ wly1
(fun (q : c = f) (s : wrz0 @ 1 = 1 @ q) =>
(fun (q0 : e = f) (s0 : wly1 @ 1 = 1 @ q0) =>
(fun (q1 : b = c) (s1 : wly0 @ 1 = 1 @ q1) =>
(fun (q2 : d = e) (s2 : wlx1 @ 1 = 1 @ q2) =>
(fun (q3 : a = b) (s3 : wlx0 @ 1 = 1 @ q3) =>
let EH_x_z :=
((rlucancel_inv (s3 [-] urnat_z1))^ @
wlrnat_x_z) @ rlucancel_inv (urnat_z2 [-] s2)
in
let EH_y_z :=
((rlucancel_inv (s1 [-] s))^ @ wlrnat_y_z) @
rlucancel_inv (urnat_z1 [-] s0) in
let EH_xy_z :=
((rlucancel_inv (...^ @ ... [-] s))^ @
((whiskerR 1 wrz0)^ @
((...) @ whiskerL wrz2 1 ))) @
rlucancel_inv
(urnat_z2 [-] (whiskerR 1 1 )^ @ (s2 [-] s0))
in
EH_xy_z @
(concat_p_pp z2 q2 q0 @ whiskerR EH_x_z^ q0) =
(concat_pp_p q3 q1 q @ whiskerL q3 EH_y_z) @
concat_p_pp q3 z1 q0) wlx0 (rlucancel 1 )) wlx1
(rlucancel 1 )) wly0 (rlucancel 1 )) wly1
(rlucancel 1 )) wrz0 (rlucancel 1 )
revert z1 urnat_z1.X : Type a, b, c, d, e, f : X wlx0 : a = b wlx1 : d = e wly0 : b = c wly1 : e = f wrz0 : c = f wrz1 : b = e wrz2, z2 : a = d urnat_z2 : wrz2 @ 1 = 1 @ z2 wlrnat_x_z : wlx0 @ wrz1 = wrz2 @ wlx1 wlrnat_y_z : wly0 @ wrz0 = wrz1 @ wly1
forall (z1 : b = e) (urnat_z1 : wrz1 @ 1 = 1 @ z1),
(fun (q : c = f) (s : wrz0 @ 1 = 1 @ q) =>
(fun (q0 : e = f) (s0 : wly1 @ 1 = 1 @ q0) =>
(fun (q1 : b = c) (s1 : wly0 @ 1 = 1 @ q1) =>
(fun (q2 : d = e) (s2 : wlx1 @ 1 = 1 @ q2) =>
(fun (q3 : a = b) (s3 : wlx0 @ 1 = 1 @ q3) =>
let EH_x_z :=
((rlucancel_inv (s3 [-] urnat_z1))^ @
wlrnat_x_z) @ rlucancel_inv (urnat_z2 [-] s2)
in
let EH_y_z :=
((rlucancel_inv (s1 [-] s))^ @ wlrnat_y_z) @
rlucancel_inv (urnat_z1 [-] s0) in
let EH_xy_z :=
((rlucancel_inv (... [-] s))^ @
((whiskerR 1 wrz0)^ @ (... @ ...))) @
rlucancel_inv
(urnat_z2 [-] (whiskerR 1 1 )^ @ (s2 [-] s0))
in
EH_xy_z @
(concat_p_pp z2 q2 q0 @ whiskerR EH_x_z^ q0) =
(concat_pp_p q3 q1 q @ whiskerL q3 EH_y_z) @
concat_p_pp q3 z1 q0) wlx0 (rlucancel 1 )) wlx1
(rlucancel 1 )) wly0 (rlucancel 1 )) wly1
(rlucancel 1 )) wrz0 (rlucancel 1 )
snapply equiv_path_ind_rlucancel. X : Type a, b, c, d, e, f : X wlx0 : a = b wlx1 : d = e wly0 : b = c wly1 : e = f wrz0 : c = f wrz1 : b = e wrz2, z2 : a = d urnat_z2 : wrz2 @ 1 = 1 @ z2 wlrnat_x_z : wlx0 @ wrz1 = wrz2 @ wlx1 wlrnat_y_z : wly0 @ wrz0 = wrz1 @ wly1
(fun (q : b = e) (s : wrz1 @ 1 = 1 @ q) =>
(fun (q0 : c = f) (s0 : wrz0 @ 1 = 1 @ q0) =>
(fun (q1 : e = f) (s1 : wly1 @ 1 = 1 @ q1) =>
(fun (q2 : b = c) (s2 : wly0 @ 1 = 1 @ q2) =>
(fun (q3 : d = e) (s3 : wlx1 @ 1 = 1 @ q3) =>
(fun (q4 : a = b) (s4 : wlx0 @ 1 = 1 @ q4) =>
let EH_x_z :=
((rlucancel_inv (s4 [-] s))^ @ wlrnat_x_z) @
rlucancel_inv (urnat_z2 [-] s3) in
let EH_y_z :=
((rlucancel_inv (...))^ @ wlrnat_y_z) @
rlucancel_inv (s [-] s1) in
let EH_xy_z :=
((rlucancel_inv ...)^ @ (...^ @ ...)) @
rlucancel_inv (urnat_z2 [-] ...^ @ ...) in
EH_xy_z @
(concat_p_pp z2 q3 q1 @ whiskerR EH_x_z^ q1) =
(concat_pp_p q4 q2 q0 @ whiskerL q4 EH_y_z) @
concat_p_pp q4 q q1) wlx0 (rlucancel 1 )) wlx1
(rlucancel 1 )) wly0 (rlucancel 1 )) wly1
(rlucancel 1 )) wrz0 (rlucancel 1 )) wrz1
(rlucancel 1 )
revert z2 urnat_z2.X : Type a, b, c, d, e, f : X wlx0 : a = b wlx1 : d = e wly0 : b = c wly1 : e = f wrz0 : c = f wrz1 : b = e wrz2 : a = d wlrnat_x_z : wlx0 @ wrz1 = wrz2 @ wlx1 wlrnat_y_z : wly0 @ wrz0 = wrz1 @ wly1
forall (z2 : a = d) (urnat_z2 : wrz2 @ 1 = 1 @ z2),
(fun (q : b = e) (s : wrz1 @ 1 = 1 @ q) =>
(fun (q0 : c = f) (s0 : wrz0 @ 1 = 1 @ q0) =>
(fun (q1 : e = f) (s1 : wly1 @ 1 = 1 @ q1) =>
(fun (q2 : b = c) (s2 : wly0 @ 1 = 1 @ q2) =>
(fun (q3 : d = e) (s3 : wlx1 @ 1 = 1 @ q3) =>
(fun (q4 : a = b) (s4 : wlx0 @ 1 = 1 @ q4) =>
let EH_x_z :=
((rlucancel_inv (...))^ @ wlrnat_x_z) @
rlucancel_inv (urnat_z2 [-] s3) in
let EH_y_z :=
((rlucancel_inv ...)^ @ wlrnat_y_z) @
rlucancel_inv (s [-] s1) in
let EH_xy_z :=
((...)^ @ (...)) @
rlucancel_inv (urnat_z2 [-] ...) in
EH_xy_z @
(concat_p_pp z2 q3 q1 @ whiskerR EH_x_z^ q1) =
(concat_pp_p q4 q2 q0 @ whiskerL q4 EH_y_z) @
concat_p_pp q4 q q1) wlx0 (rlucancel 1 )) wlx1
(rlucancel 1 )) wly0 (rlucancel 1 )) wly1
(rlucancel 1 )) wrz0 (rlucancel 1 )) wrz1
(rlucancel 1 )
snapply equiv_path_ind_rlucancel. X : Type a, b, c, d, e, f : X wlx0 : a = b wlx1 : d = e wly0 : b = c wly1 : e = f wrz0 : c = f wrz1 : b = e wrz2 : a = d wlrnat_x_z : wlx0 @ wrz1 = wrz2 @ wlx1 wlrnat_y_z : wly0 @ wrz0 = wrz1 @ wly1
(fun (q : a = d) (s : wrz2 @ 1 = 1 @ q) =>
(fun (q0 : b = e) (s0 : wrz1 @ 1 = 1 @ q0) =>
(fun (q1 : c = f) (s1 : wrz0 @ 1 = 1 @ q1) =>
(fun (q2 : e = f) (s2 : wly1 @ 1 = 1 @ q2) =>
(fun (q3 : b = c) (s3 : wly0 @ 1 = 1 @ q3) =>
(fun (q4 : d = e) (s4 : wlx1 @ 1 = 1 @ q4) =>
(fun (q5 : a = b) (s5 : wlx0 @ 1 = 1 @ q5) =>
let EH_x_z :=
((...)^ @ wlrnat_x_z) @
rlucancel_inv (s [-] s4) in
let EH_y_z :=
(...^ @ wlrnat_y_z) @ rlucancel_inv (...) in
let EH_xy_z := (...) @ rlucancel_inv ... in
EH_xy_z @ (...) = (...) @ concat_p_pp q5 q0 q2)
wlx0 (rlucancel 1 )) wlx1 (rlucancel 1 )) wly0
(rlucancel 1 )) wly1 (rlucancel 1 )) wrz0
(rlucancel 1 )) wrz1 (rlucancel 1 )) wrz2
(rlucancel 1 )
destruct wlx0, wlx1, wly0, wly1.X : Type a, d : X wlx0 : a = a wlx1 : d = d wly0 : a = a wly1 : d = d wrz0, wrz1, wrz2 : a = d wlrnat_x_z : 1 @ wrz1 = wrz2 @ 1 wlrnat_y_z : 1 @ wrz0 = wrz1 @ 1
let EH_x_z :=
((rlucancel_inv (rlucancel 1 [-] rlucancel 1 ))^ @
wlrnat_x_z) @
rlucancel_inv (rlucancel 1 [-] rlucancel 1 ) in
let EH_y_z :=
((rlucancel_inv (rlucancel 1 [-] rlucancel 1 ))^ @
wlrnat_y_z) @
rlucancel_inv (rlucancel 1 [-] rlucancel 1 ) in
let EH_xy_z :=
((rlucancel_inv
((whiskerR 1 1 )^ @ (rlucancel 1 [-] rlucancel 1 )
[-] rlucancel 1 ))^ @
((whiskerR 1 wrz0)^ @
((wlrnat_x_z [-] wlrnat_y_z) @ whiskerL wrz2 1 ))) @
rlucancel_inv
(rlucancel 1
[-] (whiskerR 1 1 )^ @
(rlucancel 1 [-] rlucancel 1 )) in
EH_xy_z @ (concat_p_pp wrz2 1 1 @ whiskerR EH_x_z^ 1 ) =
(concat_pp_p 1 1 wrz0 @ whiskerL 1 EH_y_z) @
concat_p_pp 1 wrz1 1
clear wlx0 wlx1 wly0 wly1.X : Type a, d : X wrz0, wrz1, wrz2 : a = d wlrnat_x_z : 1 @ wrz1 = wrz2 @ 1 wlrnat_y_z : 1 @ wrz0 = wrz1 @ 1
let EH_x_z :=
((rlucancel_inv (rlucancel 1 [-] rlucancel 1 ))^ @
wlrnat_x_z) @
rlucancel_inv (rlucancel 1 [-] rlucancel 1 ) in
let EH_y_z :=
((rlucancel_inv (rlucancel 1 [-] rlucancel 1 ))^ @
wlrnat_y_z) @
rlucancel_inv (rlucancel 1 [-] rlucancel 1 ) in
let EH_xy_z :=
((rlucancel_inv
((whiskerR 1 1 )^ @ (rlucancel 1 [-] rlucancel 1 )
[-] rlucancel 1 ))^ @
((whiskerR 1 wrz0)^ @
((wlrnat_x_z [-] wlrnat_y_z) @ whiskerL wrz2 1 ))) @
rlucancel_inv
(rlucancel 1
[-] (whiskerR 1 1 )^ @
(rlucancel 1 [-] rlucancel 1 )) in
EH_xy_z @ (concat_p_pp wrz2 1 1 @ whiskerR EH_x_z^ 1 ) =
(concat_pp_p 1 1 wrz0 @ whiskerL 1 EH_y_z) @
concat_p_pp 1 wrz1 1
revert wrz2 wlrnat_x_z.X : Type a, d : X wrz0, wrz1 : a = d wlrnat_y_z : 1 @ wrz0 = wrz1 @ 1
forall (wrz2 : a = d)
(wlrnat_x_z : 1 @ wrz1 = wrz2 @ 1 ),
let EH_x_z :=
((rlucancel_inv (rlucancel 1 [-] rlucancel 1 ))^ @
wlrnat_x_z) @
rlucancel_inv (rlucancel 1 [-] rlucancel 1 ) in
let EH_y_z :=
((rlucancel_inv (rlucancel 1 [-] rlucancel 1 ))^ @
wlrnat_y_z) @
rlucancel_inv (rlucancel 1 [-] rlucancel 1 ) in
let EH_xy_z :=
((rlucancel_inv
((whiskerR 1 1 )^ @ (rlucancel 1 [-] rlucancel 1 )
[-] rlucancel 1 ))^ @
((whiskerR 1 wrz0)^ @
((wlrnat_x_z [-] wlrnat_y_z) @ whiskerL wrz2 1 ))) @
rlucancel_inv
(rlucancel 1
[-] (whiskerR 1 1 )^ @
(rlucancel 1 [-] rlucancel 1 )) in
EH_xy_z @ (concat_p_pp wrz2 1 1 @ whiskerR EH_x_z^ 1 ) =
(concat_pp_p 1 1 wrz0 @ whiskerL 1 EH_y_z) @
concat_p_pp 1 wrz1 1
snapply equiv_path_ind_lrucancel. X : Type a, d : X wrz0, wrz1 : a = d wlrnat_y_z : 1 @ wrz0 = wrz1 @ 1
(fun (q : a = d) (s : 1 @ wrz1 = q @ 1 ) =>
let EH_x_z :=
((rlucancel_inv (rlucancel 1 [-] rlucancel 1 ))^ @ s) @
rlucancel_inv (rlucancel 1 [-] rlucancel 1 ) in
let EH_y_z :=
((rlucancel_inv (rlucancel 1 [-] rlucancel 1 ))^ @
wlrnat_y_z) @
rlucancel_inv (rlucancel 1 [-] rlucancel 1 ) in
let EH_xy_z :=
((rlucancel_inv
((whiskerR 1 1 )^ @
(rlucancel 1 [-] rlucancel 1 ) [-] rlucancel 1 ))^ @
((whiskerR 1 wrz0)^ @
((s [-] wlrnat_y_z) @ whiskerL q 1 ))) @
rlucancel_inv
(rlucancel 1
[-] (whiskerR 1 1 )^ @
(rlucancel 1 [-] rlucancel 1 )) in
EH_xy_z @ (concat_p_pp q 1 1 @ whiskerR EH_x_z^ 1 ) =
(concat_pp_p 1 1 wrz0 @ whiskerL 1 EH_y_z) @
concat_p_pp 1 wrz1 1 ) wrz1 (lrucancel 1 )
revert wrz1 wlrnat_y_z.X : Type a, d : X wrz0 : a = d
forall (wrz1 : a = d)
(wlrnat_y_z : 1 @ wrz0 = wrz1 @ 1 ),
(fun (q : a = d) (s : 1 @ wrz1 = q @ 1 ) =>
let EH_x_z :=
((rlucancel_inv (rlucancel 1 [-] rlucancel 1 ))^ @ s) @
rlucancel_inv (rlucancel 1 [-] rlucancel 1 ) in
let EH_y_z :=
((rlucancel_inv (rlucancel 1 [-] rlucancel 1 ))^ @
wlrnat_y_z) @
rlucancel_inv (rlucancel 1 [-] rlucancel 1 ) in
let EH_xy_z :=
((rlucancel_inv
((whiskerR 1 1 )^ @
(rlucancel 1 [-] rlucancel 1 ) [-] rlucancel 1 ))^ @
((whiskerR 1 wrz0)^ @
((s [-] wlrnat_y_z) @ whiskerL q 1 ))) @
rlucancel_inv
(rlucancel 1
[-] (whiskerR 1 1 )^ @
(rlucancel 1 [-] rlucancel 1 )) in
EH_xy_z @ (concat_p_pp q 1 1 @ whiskerR EH_x_z^ 1 ) =
(concat_pp_p 1 1 wrz0 @ whiskerL 1 EH_y_z) @
concat_p_pp 1 wrz1 1 ) wrz1 (lrucancel 1 )
snapply equiv_path_ind_lrucancel. X : Type a, d : X wrz0 : a = d
(fun (q : a = d) (s : 1 @ wrz0 = q @ 1 ) =>
(fun (q0 : a = d) (s0 : 1 @ q = q0 @ 1 ) =>
let EH_x_z :=
((rlucancel_inv (rlucancel 1 [-] rlucancel 1 ))^ @
s0) @ rlucancel_inv (rlucancel 1 [-] rlucancel 1 )
in
let EH_y_z :=
((rlucancel_inv (rlucancel 1 [-] rlucancel 1 ))^ @
s) @ rlucancel_inv (rlucancel 1 [-] rlucancel 1 )
in
let EH_xy_z :=
((rlucancel_inv
((whiskerR 1 1 )^ @
(rlucancel 1 [-] rlucancel 1 ) [-] rlucancel 1 ))^ @
((whiskerR 1 wrz0)^ @
((s0 [-] s) @ whiskerL q0 1 ))) @
rlucancel_inv
(rlucancel 1
[-] (whiskerR 1 1 )^ @
(rlucancel 1 [-] rlucancel 1 )) in
EH_xy_z @ (concat_p_pp q0 1 1 @ whiskerR EH_x_z^ 1 ) =
(concat_pp_p 1 1 wrz0 @ whiskerL 1 EH_y_z) @
concat_p_pp 1 q 1 ) q (lrucancel 1 )) wrz0
(lrucancel 1 )
destruct wrz0.X : Type a : X wrz0 : a = a
let EH_x_z :=
((rlucancel_inv (rlucancel 1 [-] rlucancel 1 ))^ @
lrucancel 1 ) @
rlucancel_inv (rlucancel 1 [-] rlucancel 1 ) in
let EH_y_z :=
((rlucancel_inv (rlucancel 1 [-] rlucancel 1 ))^ @
lrucancel 1 ) @
rlucancel_inv (rlucancel 1 [-] rlucancel 1 ) in
let EH_xy_z :=
((rlucancel_inv
((whiskerR 1 1 )^ @ (rlucancel 1 [-] rlucancel 1 )
[-] rlucancel 1 ))^ @
((whiskerR 1 1 )^ @
((lrucancel 1 [-] lrucancel 1 ) @ whiskerL 1 1 ))) @
rlucancel_inv
(rlucancel 1
[-] (whiskerR 1 1 )^ @
(rlucancel 1 [-] rlucancel 1 )) in
EH_xy_z @ (concat_p_pp 1 1 1 @ whiskerR EH_x_z^ 1 ) =
(concat_pp_p 1 1 1 @ whiskerL 1 EH_y_z) @
concat_p_pp 1 1 1
clear wrz0.X : Type a : X
let EH_x_z :=
((rlucancel_inv (rlucancel 1 [-] rlucancel 1 ))^ @
lrucancel 1 ) @
rlucancel_inv (rlucancel 1 [-] rlucancel 1 ) in
let EH_y_z :=
((rlucancel_inv (rlucancel 1 [-] rlucancel 1 ))^ @
lrucancel 1 ) @
rlucancel_inv (rlucancel 1 [-] rlucancel 1 ) in
let EH_xy_z :=
((rlucancel_inv
((whiskerR 1 1 )^ @ (rlucancel 1 [-] rlucancel 1 )
[-] rlucancel 1 ))^ @
((whiskerR 1 1 )^ @
((lrucancel 1 [-] lrucancel 1 ) @ whiskerL 1 1 ))) @
rlucancel_inv
(rlucancel 1
[-] (whiskerR 1 1 )^ @
(rlucancel 1 [-] rlucancel 1 )) in
EH_xy_z @ (concat_p_pp 1 1 1 @ whiskerR EH_x_z^ 1 ) =
(concat_pp_p 1 1 1 @ whiskerL 1 EH_y_z) @
concat_p_pp 1 1 1
reflexivity .
Defined .
End eh_pp_p .
Theorem eh_pp_p {X } {a : X} (p q r : idpath a = idpath a) :
eh (p @ q) r @ (concat_p_pp _ _ _ @ whiskerR (eh p r)^ _) =
concat_pp_p _ _ _ @ whiskerL _ (eh q r) @ concat_p_pp _ _ _.X : Type a : X p, q, r : 1 = 1
eh (p @ q) r @
(concat_p_pp r p q @ whiskerR (eh p r)^ q) =
(concat_pp_p p q r @ whiskerL p (eh q r)) @
concat_p_pp p r q
Proof .X : Type a : X p, q, r : 1 = 1
eh (p @ q) r @
(concat_p_pp r p q @ whiskerR (eh p r)^ q) =
(concat_pp_p p q r @ whiskerL p (eh q r)) @
concat_p_pp p r q
napply eh_pp_p_gen. X : Type a : X p, q, r : 1 = 1
ulnat p [-] ulnat q = whiskerR ?Goal 1 @ ulnat (p @ q)
- X : Type a : X p, q, r : 1 = 1
ulnat p [-] ulnat q = whiskerR ?Goal 1 @ ulnat (p @ q)
exact (ulnat_pp p q).
- X : Type a : X p, q, r : 1 = 1
ulnat p [-] ulnat q = whiskerR ?Goal 1 @ ulnat (p @ q)
exact (ulnat_pp p q).
- X : Type a : X p, q, r : 1 = 1
(wlrnat p r [-] wlrnat q r) @
whiskerL (whiskerR r 1 ) (whiskerL_pp 1 p q)^ =
whiskerR (whiskerL_pp 1 p q)^ (whiskerR r 1 ) @
wlrnat (p @ q) r
exact (wlrnat_pp_p p q r).
Defined .
(* Syllepsis: We now prove that "eh p q" is suitably related to "eh q p". *)
Section eh_V .
Context {X : Type }.
(* 0-paths *)
Context {a b c d : X}.
(* 1-paths *)
Context {wlx0 x0 wrx0 : a = b}.
Context {wlx1 x1 wrx1 : c = d}.
Context {wly0 y0 wry0 : b = d}.
Context {wly1 y1 wry1 : a = c}.
(* 2-paths *)
Context {ulnat_x0 : wlx0 @ 1 = 1 @ x0}.
Context {urnat_x0 : wrx0 @ 1 = 1 @ x0}.
Context {ulnat_x1 : wlx1 @ 1 = 1 @ x1}.
Context {urnat_x1 : wrx1 @ 1 = 1 @ x1}.
Context {ulnat_y0 : wly0 @ 1 = 1 @ y0}.
Context {urnat_y0 : wry0 @ 1 = 1 @ y0}.
Context {ulnat_y1 : wly1 @ 1 = 1 @ y1}.
Context {urnat_y1 : wry1 @ 1 = 1 @ y1}.
Context {ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0}.
Context {ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1}.
Context {ehrnat_y0 : wry0 @ 1 = 1 @ wly0}.
Context {ehrnat_y1 : wry1 @ 1 = 1 @ wly1}.
Context {wlrnat_x_y : wlx0 @ wry0 = wry1 @ wlx1}.
Context {wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0}.
(* 3-paths *)
Hypothesis ehlnat_1p_x0 :
(ehlnat_x0 [I] urnat_x0) @ 1 = 1 @ ulnat_x0.
Hypothesis ehlnat_1p_x1 :
(ehlnat_x1 [I] urnat_x1) @ 1 = 1 @ ulnat_x1.
Hypothesis ehrnat_p1_y0 :
(ehrnat_y0 [I] ulnat_y0) @ 1 = 1 @ urnat_y0.
Hypothesis ehrnat_p1_y1 :
(ehrnat_y1 [I] ulnat_y1) @ 1 = 1 @ urnat_y1.
Hypothesis wlrnat_V_x_y :
whiskerR wlrnat_x_y _ @ (ehrnat_y1 [-] ehlnat_x1) =
(ehlnat_x0 [-] ehrnat_y0) @ whiskerL _ wlrnat_y_x^.
(** The syllepsis *)
Definition eh_V_gen :
let EH_x_y := (rlucancel_inv (ulnat_x0 [-] urnat_y0))^ @
wlrnat_x_y @ rlucancel_inv (urnat_y1 [-] ulnat_x1) in
let EH_y_x := (rlucancel_inv (ulnat_y1 [-] urnat_x1))^ @
wlrnat_y_x @ rlucancel_inv (urnat_x0 [-] ulnat_y0) in
EH_x_y @ EH_y_x = 1 .X : Type a, b, c, d : X wlx0, x0, wrx0 : a = b wlx1, x1, wrx1 : c = d wly0, y0, wry0 : b = d wly1, y1, wry1 : a = c ulnat_x0 : wlx0 @ 1 = 1 @ x0 urnat_x0 : wrx0 @ 1 = 1 @ x0 ulnat_x1 : wlx1 @ 1 = 1 @ x1 urnat_x1 : wrx1 @ 1 = 1 @ x1 ulnat_y0 : wly0 @ 1 = 1 @ y0 urnat_y0 : wry0 @ 1 = 1 @ y0 ulnat_y1 : wly1 @ 1 = 1 @ y1 urnat_y1 : wry1 @ 1 = 1 @ y1 ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehrnat_y0 : wry0 @ 1 = 1 @ wly0 ehrnat_y1 : wry1 @ 1 = 1 @ wly1 wlrnat_x_y : wlx0 @ wry0 = wry1 @ wlx1 wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 ehlnat_1p_x0 : (ehlnat_x0 [I] urnat_x0) @ 1 =
1 @ ulnat_x0 ehlnat_1p_x1 : (ehlnat_x1 [I] urnat_x1) @ 1 =
1 @ ulnat_x1 ehrnat_p1_y0 : (ehrnat_y0 [I] ulnat_y0) @ 1 =
1 @ urnat_y0 ehrnat_p1_y1 : (ehrnat_y1 [I] ulnat_y1) @ 1 =
1 @ urnat_y1 wlrnat_V_x_y : whiskerR wlrnat_x_y 1 @
(ehrnat_y1 [-] ehlnat_x1) =
(ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^
let EH_x_y :=
((rlucancel_inv (ulnat_x0 [-] urnat_y0))^ @
wlrnat_x_y) @ rlucancel_inv (urnat_y1 [-] ulnat_x1)
in
let EH_y_x :=
((rlucancel_inv (ulnat_y1 [-] urnat_x1))^ @
wlrnat_y_x) @ rlucancel_inv (urnat_x0 [-] ulnat_y0)
in
EH_x_y @ EH_y_x = 1
Proof .X : Type a, b, c, d : X wlx0, x0, wrx0 : a = b wlx1, x1, wrx1 : c = d wly0, y0, wry0 : b = d wly1, y1, wry1 : a = c ulnat_x0 : wlx0 @ 1 = 1 @ x0 urnat_x0 : wrx0 @ 1 = 1 @ x0 ulnat_x1 : wlx1 @ 1 = 1 @ x1 urnat_x1 : wrx1 @ 1 = 1 @ x1 ulnat_y0 : wly0 @ 1 = 1 @ y0 urnat_y0 : wry0 @ 1 = 1 @ y0 ulnat_y1 : wly1 @ 1 = 1 @ y1 urnat_y1 : wry1 @ 1 = 1 @ y1 ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehrnat_y0 : wry0 @ 1 = 1 @ wly0 ehrnat_y1 : wry1 @ 1 = 1 @ wly1 wlrnat_x_y : wlx0 @ wry0 = wry1 @ wlx1 wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 ehlnat_1p_x0 : (ehlnat_x0 [I] urnat_x0) @ 1 =
1 @ ulnat_x0 ehlnat_1p_x1 : (ehlnat_x1 [I] urnat_x1) @ 1 =
1 @ ulnat_x1 ehrnat_p1_y0 : (ehrnat_y0 [I] ulnat_y0) @ 1 =
1 @ urnat_y0 ehrnat_p1_y1 : (ehrnat_y1 [I] ulnat_y1) @ 1 =
1 @ urnat_y1 wlrnat_V_x_y : whiskerR wlrnat_x_y 1 @
(ehrnat_y1 [-] ehlnat_x1) =
(ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^
let EH_x_y :=
((rlucancel_inv (ulnat_x0 [-] urnat_y0))^ @
wlrnat_x_y) @ rlucancel_inv (urnat_y1 [-] ulnat_x1)
in
let EH_y_x :=
((rlucancel_inv (ulnat_y1 [-] urnat_x1))^ @
wlrnat_y_x) @ rlucancel_inv (urnat_x0 [-] ulnat_y0)
in
EH_x_y @ EH_y_x = 1
pose (H_whiskerR_wlrnat_x_y := moveL_Mp _ _ _ (moveL_pV _ _ _ (whiskerR_p1 wlrnat_x_y))).X : Type a, b, c, d : X wlx0, x0, wrx0 : a = b wlx1, x1, wrx1 : c = d wly0, y0, wry0 : b = d wly1, y1, wry1 : a = c ulnat_x0 : wlx0 @ 1 = 1 @ x0 urnat_x0 : wrx0 @ 1 = 1 @ x0 ulnat_x1 : wlx1 @ 1 = 1 @ x1 urnat_x1 : wrx1 @ 1 = 1 @ x1 ulnat_y0 : wly0 @ 1 = 1 @ y0 urnat_y0 : wry0 @ 1 = 1 @ y0 ulnat_y1 : wly1 @ 1 = 1 @ y1 urnat_y1 : wry1 @ 1 = 1 @ y1 ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehrnat_y0 : wry0 @ 1 = 1 @ wly0 ehrnat_y1 : wry1 @ 1 = 1 @ wly1 wlrnat_x_y : wlx0 @ wry0 = wry1 @ wlx1 wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 ehlnat_1p_x0 : (ehlnat_x0 [I] urnat_x0) @ 1 =
1 @ ulnat_x0 ehlnat_1p_x1 : (ehlnat_x1 [I] urnat_x1) @ 1 =
1 @ ulnat_x1 ehrnat_p1_y0 : (ehrnat_y0 [I] ulnat_y0) @ 1 =
1 @ urnat_y0 ehrnat_p1_y1 : (ehrnat_y1 [I] ulnat_y1) @ 1 =
1 @ urnat_y1 wlrnat_V_x_y : whiskerR wlrnat_x_y 1 @
(ehrnat_y1 [-] ehlnat_x1) =
(ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^ H_whiskerR_wlrnat_x_y := moveL_Mp
(wlrnat_x_y @
(concat_p1 (wry1 @ wlx1))^)
(whiskerR wlrnat_x_y 1 )
(concat_p1 (wlx0 @ wry0))
(moveL_pV
(concat_p1 (wry1 @ wlx1))
((concat_p1 (wlx0 @ wry0))^ @
whiskerR wlrnat_x_y 1 )
wlrnat_x_y
(whiskerR_p1 wlrnat_x_y)) : whiskerR wlrnat_x_y 1 =
concat_p1 (wlx0 @ wry0) @
(wlrnat_x_y @
(concat_p1 (wry1 @ wlx1))^)
let EH_x_y :=
((rlucancel_inv (ulnat_x0 [-] urnat_y0))^ @
wlrnat_x_y) @ rlucancel_inv (urnat_y1 [-] ulnat_x1)
in
let EH_y_x :=
((rlucancel_inv (ulnat_y1 [-] urnat_x1))^ @
wlrnat_y_x) @ rlucancel_inv (urnat_x0 [-] ulnat_y0)
in
EH_x_y @ EH_y_x = 1
apply moveL_pV in wlrnat_V_x_y.X : Type a, b, c, d : X wlx0, x0, wrx0 : a = b wlx1, x1, wrx1 : c = d wly0, y0, wry0 : b = d wly1, y1, wry1 : a = c ulnat_x0 : wlx0 @ 1 = 1 @ x0 urnat_x0 : wrx0 @ 1 = 1 @ x0 ulnat_x1 : wlx1 @ 1 = 1 @ x1 urnat_x1 : wrx1 @ 1 = 1 @ x1 ulnat_y0 : wly0 @ 1 = 1 @ y0 urnat_y0 : wry0 @ 1 = 1 @ y0 ulnat_y1 : wly1 @ 1 = 1 @ y1 urnat_y1 : wry1 @ 1 = 1 @ y1 ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehrnat_y0 : wry0 @ 1 = 1 @ wly0 ehrnat_y1 : wry1 @ 1 = 1 @ wly1 wlrnat_x_y : wlx0 @ wry0 = wry1 @ wlx1 wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 ehlnat_1p_x0 : (ehlnat_x0 [I] urnat_x0) @ 1 =
1 @ ulnat_x0 ehlnat_1p_x1 : (ehlnat_x1 [I] urnat_x1) @ 1 =
1 @ ulnat_x1 ehrnat_p1_y0 : (ehrnat_y0 [I] ulnat_y0) @ 1 =
1 @ urnat_y0 ehrnat_p1_y1 : (ehrnat_y1 [I] ulnat_y1) @ 1 =
1 @ urnat_y1 wlrnat_V_x_y : whiskerR wlrnat_x_y 1 =
((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^ H_whiskerR_wlrnat_x_y := moveL_Mp
(wlrnat_x_y @
(concat_p1 (wry1 @ wlx1))^)
(whiskerR wlrnat_x_y 1 )
(concat_p1 (wlx0 @ wry0))
(moveL_pV
(concat_p1 (wry1 @ wlx1))
((concat_p1 (wlx0 @ wry0))^ @
whiskerR wlrnat_x_y 1 )
wlrnat_x_y
(whiskerR_p1 wlrnat_x_y)) : whiskerR wlrnat_x_y 1 =
concat_p1 (wlx0 @ wry0) @
(wlrnat_x_y @
(concat_p1 (wry1 @ wlx1))^)
let EH_x_y :=
((rlucancel_inv (ulnat_x0 [-] urnat_y0))^ @
wlrnat_x_y) @ rlucancel_inv (urnat_y1 [-] ulnat_x1)
in
let EH_y_x :=
((rlucancel_inv (ulnat_y1 [-] urnat_x1))^ @
wlrnat_y_x) @ rlucancel_inv (urnat_x0 [-] ulnat_y0)
in
EH_x_y @ EH_y_x = 1
apply (concat H_whiskerR_wlrnat_x_y^) in wlrnat_V_x_y.X : Type a, b, c, d : X wlx0, x0, wrx0 : a = b wlx1, x1, wrx1 : c = d wly0, y0, wry0 : b = d wly1, y1, wry1 : a = c ulnat_x0 : wlx0 @ 1 = 1 @ x0 urnat_x0 : wrx0 @ 1 = 1 @ x0 ulnat_x1 : wlx1 @ 1 = 1 @ x1 urnat_x1 : wrx1 @ 1 = 1 @ x1 ulnat_y0 : wly0 @ 1 = 1 @ y0 urnat_y0 : wry0 @ 1 = 1 @ y0 ulnat_y1 : wly1 @ 1 = 1 @ y1 urnat_y1 : wry1 @ 1 = 1 @ y1 ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehrnat_y0 : wry0 @ 1 = 1 @ wly0 ehrnat_y1 : wry1 @ 1 = 1 @ wly1 wlrnat_x_y : wlx0 @ wry0 = wry1 @ wlx1 wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 ehlnat_1p_x0 : (ehlnat_x0 [I] urnat_x0) @ 1 =
1 @ ulnat_x0 ehlnat_1p_x1 : (ehlnat_x1 [I] urnat_x1) @ 1 =
1 @ ulnat_x1 ehrnat_p1_y0 : (ehrnat_y0 [I] ulnat_y0) @ 1 =
1 @ urnat_y0 ehrnat_p1_y1 : (ehrnat_y1 [I] ulnat_y1) @ 1 =
1 @ urnat_y1 wlrnat_V_x_y : concat_p1 (wlx0 @ wry0) @
(wlrnat_x_y @
(concat_p1 (wry1 @ wlx1))^) =
((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^ H_whiskerR_wlrnat_x_y := moveL_Mp
(wlrnat_x_y @
(concat_p1 (wry1 @ wlx1))^)
(whiskerR wlrnat_x_y 1 )
(concat_p1 (wlx0 @ wry0))
(moveL_pV
(concat_p1 (wry1 @ wlx1))
((concat_p1 (wlx0 @ wry0))^ @
whiskerR wlrnat_x_y 1 )
wlrnat_x_y
(whiskerR_p1 wlrnat_x_y)) : whiskerR wlrnat_x_y 1 =
concat_p1 (wlx0 @ wry0) @
(wlrnat_x_y @
(concat_p1 (wry1 @ wlx1))^)
let EH_x_y :=
((rlucancel_inv (ulnat_x0 [-] urnat_y0))^ @
wlrnat_x_y) @ rlucancel_inv (urnat_y1 [-] ulnat_x1)
in
let EH_y_x :=
((rlucancel_inv (ulnat_y1 [-] urnat_x1))^ @
wlrnat_y_x) @ rlucancel_inv (urnat_x0 [-] ulnat_y0)
in
EH_x_y @ EH_y_x = 1
apply moveL_Vp, moveL_pV in wlrnat_V_x_y.X : Type a, b, c, d : X wlx0, x0, wrx0 : a = b wlx1, x1, wrx1 : c = d wly0, y0, wry0 : b = d wly1, y1, wry1 : a = c ulnat_x0 : wlx0 @ 1 = 1 @ x0 urnat_x0 : wrx0 @ 1 = 1 @ x0 ulnat_x1 : wlx1 @ 1 = 1 @ x1 urnat_x1 : wrx1 @ 1 = 1 @ x1 ulnat_y0 : wly0 @ 1 = 1 @ y0 urnat_y0 : wry0 @ 1 = 1 @ y0 ulnat_y1 : wly1 @ 1 = 1 @ y1 urnat_y1 : wry1 @ 1 = 1 @ y1 ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehrnat_y0 : wry0 @ 1 = 1 @ wly0 ehrnat_y1 : wry1 @ 1 = 1 @ wly1 wlrnat_x_y : wlx0 @ wry0 = wry1 @ wlx1 wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 ehlnat_1p_x0 : (ehlnat_x0 [I] urnat_x0) @ 1 =
1 @ ulnat_x0 ehlnat_1p_x1 : (ehlnat_x1 [I] urnat_x1) @ 1 =
1 @ ulnat_x1 ehrnat_p1_y0 : (ehrnat_y0 [I] ulnat_y0) @ 1 =
1 @ urnat_y0 ehrnat_p1_y1 : (ehrnat_y1 [I] ulnat_y1) @ 1 =
1 @ urnat_y1 wlrnat_V_x_y : wlrnat_x_y =
((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^ H_whiskerR_wlrnat_x_y := moveL_Mp
(wlrnat_x_y @
(concat_p1 (wry1 @ wlx1))^)
(whiskerR wlrnat_x_y 1 )
(concat_p1 (wlx0 @ wry0))
(moveL_pV
(concat_p1 (wry1 @ wlx1))
((concat_p1 (wlx0 @ wry0))^ @
whiskerR wlrnat_x_y 1 )
wlrnat_x_y
(whiskerR_p1 wlrnat_x_y)) : whiskerR wlrnat_x_y 1 =
concat_p1 (wlx0 @ wry0) @
(wlrnat_x_y @
(concat_p1 (wry1 @ wlx1))^)
let EH_x_y :=
((rlucancel_inv (ulnat_x0 [-] urnat_y0))^ @
wlrnat_x_y) @ rlucancel_inv (urnat_y1 [-] ulnat_x1)
in
let EH_y_x :=
((rlucancel_inv (ulnat_y1 [-] urnat_x1))^ @
wlrnat_y_x) @ rlucancel_inv (urnat_x0 [-] ulnat_y0)
in
EH_x_y @ EH_y_x = 1
apply symmetry in wlrnat_V_x_y.X : Type a, b, c, d : X wlx0, x0, wrx0 : a = b wlx1, x1, wrx1 : c = d wly0, y0, wry0 : b = d wly1, y1, wry1 : a = c ulnat_x0 : wlx0 @ 1 = 1 @ x0 urnat_x0 : wrx0 @ 1 = 1 @ x0 ulnat_x1 : wlx1 @ 1 = 1 @ x1 urnat_x1 : wrx1 @ 1 = 1 @ x1 ulnat_y0 : wly0 @ 1 = 1 @ y0 urnat_y0 : wry0 @ 1 = 1 @ y0 ulnat_y1 : wly1 @ 1 = 1 @ y1 urnat_y1 : wry1 @ 1 = 1 @ y1 ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehrnat_y0 : wry0 @ 1 = 1 @ wly0 ehrnat_y1 : wry1 @ 1 = 1 @ wly1 wlrnat_x_y : wlx0 @ wry0 = wry1 @ wlx1 wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 ehlnat_1p_x0 : (ehlnat_x0 [I] urnat_x0) @ 1 =
1 @ ulnat_x0 ehlnat_1p_x1 : (ehlnat_x1 [I] urnat_x1) @ 1 =
1 @ ulnat_x1 ehrnat_p1_y0 : (ehrnat_y0 [I] ulnat_y0) @ 1 =
1 @ urnat_y0 ehrnat_p1_y1 : (ehrnat_y1 [I] ulnat_y1) @ 1 =
1 @ urnat_y1 wlrnat_V_x_y : ((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^ =
wlrnat_x_y H_whiskerR_wlrnat_x_y := moveL_Mp
(wlrnat_x_y @
(concat_p1 (wry1 @ wlx1))^)
(whiskerR wlrnat_x_y 1 )
(concat_p1 (wlx0 @ wry0))
(moveL_pV
(concat_p1 (wry1 @ wlx1))
((concat_p1 (wlx0 @ wry0))^ @
whiskerR wlrnat_x_y 1 )
wlrnat_x_y
(whiskerR_p1 wlrnat_x_y)) : whiskerR wlrnat_x_y 1 =
concat_p1 (wlx0 @ wry0) @
(wlrnat_x_y @
(concat_p1 (wry1 @ wlx1))^)
let EH_x_y :=
((rlucancel_inv (ulnat_x0 [-] urnat_y0))^ @
wlrnat_x_y) @ rlucancel_inv (urnat_y1 [-] ulnat_x1)
in
let EH_y_x :=
((rlucancel_inv (ulnat_y1 [-] urnat_x1))^ @
wlrnat_y_x) @ rlucancel_inv (urnat_x0 [-] ulnat_y0)
in
EH_x_y @ EH_y_x = 1
destruct wlrnat_V_x_y.X : Type a, b, c, d : X wlx0, x0, wrx0 : a = b wlx1, x1, wrx1 : c = d wly0, y0, wry0 : b = d wly1, y1, wry1 : a = c ulnat_x0 : wlx0 @ 1 = 1 @ x0 urnat_x0 : wrx0 @ 1 = 1 @ x0 ulnat_x1 : wlx1 @ 1 = 1 @ x1 urnat_x1 : wrx1 @ 1 = 1 @ x1 ulnat_y0 : wly0 @ 1 = 1 @ y0 urnat_y0 : wry0 @ 1 = 1 @ y0 ulnat_y1 : wly1 @ 1 = 1 @ y1 urnat_y1 : wry1 @ 1 = 1 @ y1 ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehrnat_y0 : wry0 @ 1 = 1 @ wly0 ehrnat_y1 : wry1 @ 1 = 1 @ wly1 wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 ehlnat_1p_x0 : (ehlnat_x0 [I] urnat_x0) @ 1 =
1 @ ulnat_x0 ehlnat_1p_x1 : (ehlnat_x1 [I] urnat_x1) @ 1 =
1 @ ulnat_x1 ehrnat_p1_y0 : (ehrnat_y0 [I] ulnat_y0) @ 1 =
1 @ urnat_y0 ehrnat_p1_y1 : (ehrnat_y1 [I] ulnat_y1) @ 1 =
1 @ urnat_y1 wlrnat_V_x_y : ((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^ =
((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^ H_whiskerR_wlrnat_x_y := moveL_Mp
((((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^) @
(concat_p1 (wry1 @ wlx1))^)
(whiskerR
(((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0
[-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1
[-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^)
1 )
(concat_p1 (wlx0 @ wry0))
(moveL_pV
(concat_p1 (wry1 @ wlx1))
((concat_p1 (wlx0 @ wry0))^ @
whiskerR
(((concat_p1
(wlx0 @ wry0))^ @
(((ehlnat_x0
[-] ehrnat_y0) @
whiskerL 1
wlrnat_y_x^) @
(ehrnat_y1
[-] ehlnat_x1)^)) @
((concat_p1
(wry1 @ wlx1))^)^)
1 )
(((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0
[-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1
[-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^)
(whiskerR_p1
(((concat_p1
(wlx0 @ wry0))^ @
(((ehlnat_x0
[-] ehrnat_y0) @
whiskerL 1
wlrnat_y_x^) @
(ehrnat_y1
[-] ehlnat_x1)^)) @
((concat_p1
(wry1 @ wlx1))^)^))) : whiskerR
(((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^)
1 =
concat_p1 (wlx0 @ wry0) @
((((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^) @
(concat_p1 (wry1 @ wlx1))^)
let EH_x_y :=
((rlucancel_inv (ulnat_x0 [-] urnat_y0))^ @
(((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^)) @
rlucancel_inv (urnat_y1 [-] ulnat_x1) in
let EH_y_x :=
((rlucancel_inv (ulnat_y1 [-] urnat_x1))^ @
wlrnat_y_x) @ rlucancel_inv (urnat_x0 [-] ulnat_y0)
in
EH_x_y @ EH_y_x = 1
clear wlrnat_V_x_y.X : Type a, b, c, d : X wlx0, x0, wrx0 : a = b wlx1, x1, wrx1 : c = d wly0, y0, wry0 : b = d wly1, y1, wry1 : a = c ulnat_x0 : wlx0 @ 1 = 1 @ x0 urnat_x0 : wrx0 @ 1 = 1 @ x0 ulnat_x1 : wlx1 @ 1 = 1 @ x1 urnat_x1 : wrx1 @ 1 = 1 @ x1 ulnat_y0 : wly0 @ 1 = 1 @ y0 urnat_y0 : wry0 @ 1 = 1 @ y0 ulnat_y1 : wly1 @ 1 = 1 @ y1 urnat_y1 : wry1 @ 1 = 1 @ y1 ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehrnat_y0 : wry0 @ 1 = 1 @ wly0 ehrnat_y1 : wry1 @ 1 = 1 @ wly1 wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 ehlnat_1p_x0 : (ehlnat_x0 [I] urnat_x0) @ 1 =
1 @ ulnat_x0 ehlnat_1p_x1 : (ehlnat_x1 [I] urnat_x1) @ 1 =
1 @ ulnat_x1 ehrnat_p1_y0 : (ehrnat_y0 [I] ulnat_y0) @ 1 =
1 @ urnat_y0 ehrnat_p1_y1 : (ehrnat_y1 [I] ulnat_y1) @ 1 =
1 @ urnat_y1 H_whiskerR_wlrnat_x_y := moveL_Mp
((((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^) @
(concat_p1 (wry1 @ wlx1))^)
(whiskerR
(((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0
[-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1
[-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^)
1 )
(concat_p1 (wlx0 @ wry0))
(moveL_pV
(concat_p1 (wry1 @ wlx1))
((concat_p1 (wlx0 @ wry0))^ @
whiskerR
(((concat_p1
(wlx0 @ wry0))^ @
(((ehlnat_x0
[-] ehrnat_y0) @
whiskerL 1
wlrnat_y_x^) @
(ehrnat_y1
[-] ehlnat_x1)^)) @
((concat_p1
(wry1 @ wlx1))^)^)
1 )
(((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0
[-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1
[-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^)
(whiskerR_p1
(((concat_p1
(wlx0 @ wry0))^ @
(((ehlnat_x0
[-] ehrnat_y0) @
whiskerL 1
wlrnat_y_x^) @
(ehrnat_y1
[-] ehlnat_x1)^)) @
((concat_p1
(wry1 @ wlx1))^)^))) : whiskerR
(((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^)
1 =
concat_p1 (wlx0 @ wry0) @
((((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^) @
(concat_p1 (wry1 @ wlx1))^)
let EH_x_y :=
((rlucancel_inv (ulnat_x0 [-] urnat_y0))^ @
(((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^)) @
rlucancel_inv (urnat_y1 [-] ulnat_x1) in
let EH_y_x :=
((rlucancel_inv (ulnat_y1 [-] urnat_x1))^ @
wlrnat_y_x) @ rlucancel_inv (urnat_x0 [-] ulnat_y0)
in
EH_x_y @ EH_y_x = 1
clear H_whiskerR_wlrnat_x_y.X : Type a, b, c, d : X wlx0, x0, wrx0 : a = b wlx1, x1, wrx1 : c = d wly0, y0, wry0 : b = d wly1, y1, wry1 : a = c ulnat_x0 : wlx0 @ 1 = 1 @ x0 urnat_x0 : wrx0 @ 1 = 1 @ x0 ulnat_x1 : wlx1 @ 1 = 1 @ x1 urnat_x1 : wrx1 @ 1 = 1 @ x1 ulnat_y0 : wly0 @ 1 = 1 @ y0 urnat_y0 : wry0 @ 1 = 1 @ y0 ulnat_y1 : wly1 @ 1 = 1 @ y1 urnat_y1 : wry1 @ 1 = 1 @ y1 ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehrnat_y0 : wry0 @ 1 = 1 @ wly0 ehrnat_y1 : wry1 @ 1 = 1 @ wly1 wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 ehlnat_1p_x0 : (ehlnat_x0 [I] urnat_x0) @ 1 =
1 @ ulnat_x0 ehlnat_1p_x1 : (ehlnat_x1 [I] urnat_x1) @ 1 =
1 @ ulnat_x1 ehrnat_p1_y0 : (ehrnat_y0 [I] ulnat_y0) @ 1 =
1 @ urnat_y0 ehrnat_p1_y1 : (ehrnat_y1 [I] ulnat_y1) @ 1 =
1 @ urnat_y1
let EH_x_y :=
((rlucancel_inv (ulnat_x0 [-] urnat_y0))^ @
(((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^)) @
rlucancel_inv (urnat_y1 [-] ulnat_x1) in
let EH_y_x :=
((rlucancel_inv (ulnat_y1 [-] urnat_x1))^ @
wlrnat_y_x) @ rlucancel_inv (urnat_x0 [-] ulnat_y0)
in
EH_x_y @ EH_y_x = 1
revert ulnat_x0 ehlnat_1p_x0.X : Type a, b, c, d : X wlx0, x0, wrx0 : a = b wlx1, x1, wrx1 : c = d wly0, y0, wry0 : b = d wly1, y1, wry1 : a = c urnat_x0 : wrx0 @ 1 = 1 @ x0 ulnat_x1 : wlx1 @ 1 = 1 @ x1 urnat_x1 : wrx1 @ 1 = 1 @ x1 ulnat_y0 : wly0 @ 1 = 1 @ y0 urnat_y0 : wry0 @ 1 = 1 @ y0 ulnat_y1 : wly1 @ 1 = 1 @ y1 urnat_y1 : wry1 @ 1 = 1 @ y1 ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehrnat_y0 : wry0 @ 1 = 1 @ wly0 ehrnat_y1 : wry1 @ 1 = 1 @ wly1 wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 ehlnat_1p_x1 : (ehlnat_x1 [I] urnat_x1) @ 1 =
1 @ ulnat_x1 ehrnat_p1_y0 : (ehrnat_y0 [I] ulnat_y0) @ 1 =
1 @ urnat_y0 ehrnat_p1_y1 : (ehrnat_y1 [I] ulnat_y1) @ 1 =
1 @ urnat_y1
forall ulnat_x0 : wlx0 @ 1 = 1 @ x0,
(ehlnat_x0 [I] urnat_x0) @ 1 = 1 @ ulnat_x0 ->
let EH_x_y :=
((rlucancel_inv (ulnat_x0 [-] urnat_y0))^ @
(((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^)) @
rlucancel_inv (urnat_y1 [-] ulnat_x1) in
let EH_y_x :=
((rlucancel_inv (ulnat_y1 [-] urnat_x1))^ @
wlrnat_y_x) @ rlucancel_inv (urnat_x0 [-] ulnat_y0)
in
EH_x_y @ EH_y_x = 1
snapply equiv_path_ind_rlucancel. X : Type a, b, c, d : X wlx0, x0, wrx0 : a = b wlx1, x1, wrx1 : c = d wly0, y0, wry0 : b = d wly1, y1, wry1 : a = c urnat_x0 : wrx0 @ 1 = 1 @ x0 ulnat_x1 : wlx1 @ 1 = 1 @ x1 urnat_x1 : wrx1 @ 1 = 1 @ x1 ulnat_y0 : wly0 @ 1 = 1 @ y0 urnat_y0 : wry0 @ 1 = 1 @ y0 ulnat_y1 : wly1 @ 1 = 1 @ y1 urnat_y1 : wry1 @ 1 = 1 @ y1 ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehrnat_y0 : wry0 @ 1 = 1 @ wly0 ehrnat_y1 : wry1 @ 1 = 1 @ wly1 wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 ehlnat_1p_x1 : (ehlnat_x1 [I] urnat_x1) @ 1 =
1 @ ulnat_x1 ehrnat_p1_y0 : (ehrnat_y0 [I] ulnat_y0) @ 1 =
1 @ urnat_y0 ehrnat_p1_y1 : (ehrnat_y1 [I] ulnat_y1) @ 1 =
1 @ urnat_y1
(fun (q : wlx0 @ 1 = 1 @ x0)
(_ : (ehlnat_x0 [I] urnat_x0) @ 1 = 1 @ q) =>
let EH_x_y :=
((rlucancel_inv (q [-] urnat_y0))^ @
(((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^)) @
rlucancel_inv (urnat_y1 [-] ulnat_x1) in
let EH_y_x :=
((rlucancel_inv (ulnat_y1 [-] urnat_x1))^ @
wlrnat_y_x) @
rlucancel_inv (urnat_x0 [-] ulnat_y0) in
EH_x_y @ EH_y_x = 1 ) (ehlnat_x0 [I] urnat_x0)
(rlucancel 1 )
revert ulnat_x1 ehlnat_1p_x1.X : Type a, b, c, d : X wlx0, x0, wrx0 : a = b wlx1, x1, wrx1 : c = d wly0, y0, wry0 : b = d wly1, y1, wry1 : a = c urnat_x0 : wrx0 @ 1 = 1 @ x0 urnat_x1 : wrx1 @ 1 = 1 @ x1 ulnat_y0 : wly0 @ 1 = 1 @ y0 urnat_y0 : wry0 @ 1 = 1 @ y0 ulnat_y1 : wly1 @ 1 = 1 @ y1 urnat_y1 : wry1 @ 1 = 1 @ y1 ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehrnat_y0 : wry0 @ 1 = 1 @ wly0 ehrnat_y1 : wry1 @ 1 = 1 @ wly1 wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 ehrnat_p1_y0 : (ehrnat_y0 [I] ulnat_y0) @ 1 =
1 @ urnat_y0 ehrnat_p1_y1 : (ehrnat_y1 [I] ulnat_y1) @ 1 =
1 @ urnat_y1
forall ulnat_x1 : wlx1 @ 1 = 1 @ x1,
(ehlnat_x1 [I] urnat_x1) @ 1 = 1 @ ulnat_x1 ->
(fun (q : wlx0 @ 1 = 1 @ x0)
(_ : (ehlnat_x0 [I] urnat_x0) @ 1 = 1 @ q) =>
let EH_x_y :=
((rlucancel_inv (q [-] urnat_y0))^ @
(((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^)) @
rlucancel_inv (urnat_y1 [-] ulnat_x1) in
let EH_y_x :=
((rlucancel_inv (ulnat_y1 [-] urnat_x1))^ @
wlrnat_y_x) @
rlucancel_inv (urnat_x0 [-] ulnat_y0) in
EH_x_y @ EH_y_x = 1 ) (ehlnat_x0 [I] urnat_x0)
(rlucancel 1 )
snapply equiv_path_ind_rlucancel. X : Type a, b, c, d : X wlx0, x0, wrx0 : a = b wlx1, x1, wrx1 : c = d wly0, y0, wry0 : b = d wly1, y1, wry1 : a = c urnat_x0 : wrx0 @ 1 = 1 @ x0 urnat_x1 : wrx1 @ 1 = 1 @ x1 ulnat_y0 : wly0 @ 1 = 1 @ y0 urnat_y0 : wry0 @ 1 = 1 @ y0 ulnat_y1 : wly1 @ 1 = 1 @ y1 urnat_y1 : wry1 @ 1 = 1 @ y1 ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehrnat_y0 : wry0 @ 1 = 1 @ wly0 ehrnat_y1 : wry1 @ 1 = 1 @ wly1 wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 ehrnat_p1_y0 : (ehrnat_y0 [I] ulnat_y0) @ 1 =
1 @ urnat_y0 ehrnat_p1_y1 : (ehrnat_y1 [I] ulnat_y1) @ 1 =
1 @ urnat_y1
(fun (q : wlx1 @ 1 = 1 @ x1)
(_ : (ehlnat_x1 [I] urnat_x1) @ 1 = 1 @ q) =>
(fun (q0 : wlx0 @ 1 = 1 @ x0)
(_ : (ehlnat_x0 [I] urnat_x0) @ 1 = 1 @ q0) =>
let EH_x_y :=
((rlucancel_inv (q0 [-] urnat_y0))^ @
(((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^)) @
rlucancel_inv (urnat_y1 [-] q) in
let EH_y_x :=
((rlucancel_inv (ulnat_y1 [-] urnat_x1))^ @
wlrnat_y_x) @
rlucancel_inv (urnat_x0 [-] ulnat_y0) in
EH_x_y @ EH_y_x = 1 ) (ehlnat_x0 [I] urnat_x0)
(rlucancel 1 )) (ehlnat_x1 [I] urnat_x1)
(rlucancel 1 )
revert urnat_y0 ehrnat_p1_y0.X : Type a, b, c, d : X wlx0, x0, wrx0 : a = b wlx1, x1, wrx1 : c = d wly0, y0, wry0 : b = d wly1, y1, wry1 : a = c urnat_x0 : wrx0 @ 1 = 1 @ x0 urnat_x1 : wrx1 @ 1 = 1 @ x1 ulnat_y0 : wly0 @ 1 = 1 @ y0 ulnat_y1 : wly1 @ 1 = 1 @ y1 urnat_y1 : wry1 @ 1 = 1 @ y1 ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehrnat_y0 : wry0 @ 1 = 1 @ wly0 ehrnat_y1 : wry1 @ 1 = 1 @ wly1 wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 ehrnat_p1_y1 : (ehrnat_y1 [I] ulnat_y1) @ 1 =
1 @ urnat_y1
forall urnat_y0 : wry0 @ 1 = 1 @ y0,
(ehrnat_y0 [I] ulnat_y0) @ 1 = 1 @ urnat_y0 ->
(fun (q : wlx1 @ 1 = 1 @ x1)
(_ : (ehlnat_x1 [I] urnat_x1) @ 1 = 1 @ q) =>
(fun (q0 : wlx0 @ 1 = 1 @ x0)
(_ : (ehlnat_x0 [I] urnat_x0) @ 1 = 1 @ q0) =>
let EH_x_y :=
((rlucancel_inv (q0 [-] urnat_y0))^ @
(((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^)) @
rlucancel_inv (urnat_y1 [-] q) in
let EH_y_x :=
((rlucancel_inv (ulnat_y1 [-] urnat_x1))^ @
wlrnat_y_x) @
rlucancel_inv (urnat_x0 [-] ulnat_y0) in
EH_x_y @ EH_y_x = 1 ) (ehlnat_x0 [I] urnat_x0)
(rlucancel 1 )) (ehlnat_x1 [I] urnat_x1)
(rlucancel 1 )
snapply equiv_path_ind_rlucancel. X : Type a, b, c, d : X wlx0, x0, wrx0 : a = b wlx1, x1, wrx1 : c = d wly0, y0, wry0 : b = d wly1, y1, wry1 : a = c urnat_x0 : wrx0 @ 1 = 1 @ x0 urnat_x1 : wrx1 @ 1 = 1 @ x1 ulnat_y0 : wly0 @ 1 = 1 @ y0 ulnat_y1 : wly1 @ 1 = 1 @ y1 urnat_y1 : wry1 @ 1 = 1 @ y1 ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehrnat_y0 : wry0 @ 1 = 1 @ wly0 ehrnat_y1 : wry1 @ 1 = 1 @ wly1 wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 ehrnat_p1_y1 : (ehrnat_y1 [I] ulnat_y1) @ 1 =
1 @ urnat_y1
(fun (q : wry0 @ 1 = 1 @ y0)
(_ : (ehrnat_y0 [I] ulnat_y0) @ 1 = 1 @ q) =>
(fun (q0 : wlx1 @ 1 = 1 @ x1)
(_ : (ehlnat_x1 [I] urnat_x1) @ 1 = 1 @ q0) =>
(fun (q1 : wlx0 @ 1 = 1 @ x0)
(_ : (ehlnat_x0 [I] urnat_x0) @ 1 = 1 @ q1) =>
let EH_x_y :=
((rlucancel_inv (q1 [-] q))^ @
(((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^)) @
rlucancel_inv (urnat_y1 [-] q0) in
let EH_y_x :=
((rlucancel_inv (ulnat_y1 [-] urnat_x1))^ @
wlrnat_y_x) @
rlucancel_inv (urnat_x0 [-] ulnat_y0) in
EH_x_y @ EH_y_x = 1 ) (ehlnat_x0 [I] urnat_x0)
(rlucancel 1 )) (ehlnat_x1 [I] urnat_x1)
(rlucancel 1 )) (ehrnat_y0 [I] ulnat_y0)
(rlucancel 1 )
revert urnat_y1 ehrnat_p1_y1.X : Type a, b, c, d : X wlx0, x0, wrx0 : a = b wlx1, x1, wrx1 : c = d wly0, y0, wry0 : b = d wly1, y1, wry1 : a = c urnat_x0 : wrx0 @ 1 = 1 @ x0 urnat_x1 : wrx1 @ 1 = 1 @ x1 ulnat_y0 : wly0 @ 1 = 1 @ y0 ulnat_y1 : wly1 @ 1 = 1 @ y1 ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehrnat_y0 : wry0 @ 1 = 1 @ wly0 ehrnat_y1 : wry1 @ 1 = 1 @ wly1 wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0
forall urnat_y1 : wry1 @ 1 = 1 @ y1,
(ehrnat_y1 [I] ulnat_y1) @ 1 = 1 @ urnat_y1 ->
(fun (q : wry0 @ 1 = 1 @ y0)
(_ : (ehrnat_y0 [I] ulnat_y0) @ 1 = 1 @ q) =>
(fun (q0 : wlx1 @ 1 = 1 @ x1)
(_ : (ehlnat_x1 [I] urnat_x1) @ 1 = 1 @ q0) =>
(fun (q1 : wlx0 @ 1 = 1 @ x0)
(_ : (ehlnat_x0 [I] urnat_x0) @ 1 = 1 @ q1) =>
let EH_x_y :=
((rlucancel_inv (q1 [-] q))^ @
(((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^)) @
rlucancel_inv (urnat_y1 [-] q0) in
let EH_y_x :=
((rlucancel_inv (ulnat_y1 [-] urnat_x1))^ @
wlrnat_y_x) @
rlucancel_inv (urnat_x0 [-] ulnat_y0) in
EH_x_y @ EH_y_x = 1 ) (ehlnat_x0 [I] urnat_x0)
(rlucancel 1 )) (ehlnat_x1 [I] urnat_x1)
(rlucancel 1 )) (ehrnat_y0 [I] ulnat_y0)
(rlucancel 1 )
snapply equiv_path_ind_rlucancel. X : Type a, b, c, d : X wlx0, x0, wrx0 : a = b wlx1, x1, wrx1 : c = d wly0, y0, wry0 : b = d wly1, y1, wry1 : a = c urnat_x0 : wrx0 @ 1 = 1 @ x0 urnat_x1 : wrx1 @ 1 = 1 @ x1 ulnat_y0 : wly0 @ 1 = 1 @ y0 ulnat_y1 : wly1 @ 1 = 1 @ y1 ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehrnat_y0 : wry0 @ 1 = 1 @ wly0 ehrnat_y1 : wry1 @ 1 = 1 @ wly1 wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0
(fun (q : wry1 @ 1 = 1 @ y1)
(_ : (ehrnat_y1 [I] ulnat_y1) @ 1 = 1 @ q) =>
(fun (q0 : wry0 @ 1 = 1 @ y0)
(_ : (ehrnat_y0 [I] ulnat_y0) @ 1 = 1 @ q0) =>
(fun (q1 : wlx1 @ 1 = 1 @ x1)
(_ : (ehlnat_x1 [I] urnat_x1) @ 1 = 1 @ q1) =>
(fun (q2 : wlx0 @ 1 = 1 @ x0)
(_ : (ehlnat_x0 [I] urnat_x0) @ 1 = 1 @ q2) =>
let EH_x_y :=
((rlucancel_inv (q2 [-] q0))^ @
(((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^)) @
rlucancel_inv (q [-] q1) in
let EH_y_x :=
((rlucancel_inv (ulnat_y1 [-] urnat_x1))^ @
wlrnat_y_x) @
rlucancel_inv (urnat_x0 [-] ulnat_y0) in
EH_x_y @ EH_y_x = 1 ) (ehlnat_x0 [I] urnat_x0)
(rlucancel 1 )) (ehlnat_x1 [I] urnat_x1)
(rlucancel 1 )) (ehrnat_y0 [I] ulnat_y0)
(rlucancel 1 )) (ehrnat_y1 [I] ulnat_y1)
(rlucancel 1 )
revert x0 urnat_x0.X : Type a, b, c, d : X wlx0, wrx0 : a = b wlx1, x1, wrx1 : c = d wly0, y0, wry0 : b = d wly1, y1, wry1 : a = c urnat_x1 : wrx1 @ 1 = 1 @ x1 ulnat_y0 : wly0 @ 1 = 1 @ y0 ulnat_y1 : wly1 @ 1 = 1 @ y1 ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehrnat_y0 : wry0 @ 1 = 1 @ wly0 ehrnat_y1 : wry1 @ 1 = 1 @ wly1 wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0
forall (x0 : a = b) (urnat_x0 : wrx0 @ 1 = 1 @ x0),
(fun (q : wry1 @ 1 = 1 @ y1)
(_ : (ehrnat_y1 [I] ulnat_y1) @ 1 = 1 @ q) =>
(fun (q0 : wry0 @ 1 = 1 @ y0)
(_ : (ehrnat_y0 [I] ulnat_y0) @ 1 = 1 @ q0) =>
(fun (q1 : wlx1 @ 1 = 1 @ x1)
(_ : (ehlnat_x1 [I] urnat_x1) @ 1 = 1 @ q1) =>
(fun (q2 : wlx0 @ 1 = 1 @ x0)
(_ : (ehlnat_x0 [I] urnat_x0) @ 1 = 1 @ q2) =>
let EH_x_y :=
((rlucancel_inv (q2 [-] q0))^ @
(((concat_p1 (wlx0 @ wry0))^ @
(((...) @ whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^)) @
rlucancel_inv (q [-] q1) in
let EH_y_x :=
((rlucancel_inv (ulnat_y1 [-] urnat_x1))^ @
wlrnat_y_x) @
rlucancel_inv (urnat_x0 [-] ulnat_y0) in
EH_x_y @ EH_y_x = 1 ) (ehlnat_x0 [I] urnat_x0)
(rlucancel 1 )) (ehlnat_x1 [I] urnat_x1)
(rlucancel 1 )) (ehrnat_y0 [I] ulnat_y0)
(rlucancel 1 )) (ehrnat_y1 [I] ulnat_y1)
(rlucancel 1 )
snapply equiv_path_ind_rlucancel. X : Type a, b, c, d : X wlx0, wrx0 : a = b wlx1, x1, wrx1 : c = d wly0, y0, wry0 : b = d wly1, y1, wry1 : a = c urnat_x1 : wrx1 @ 1 = 1 @ x1 ulnat_y0 : wly0 @ 1 = 1 @ y0 ulnat_y1 : wly1 @ 1 = 1 @ y1 ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehrnat_y0 : wry0 @ 1 = 1 @ wly0 ehrnat_y1 : wry1 @ 1 = 1 @ wly1 wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0
(fun (q : a = b) (s : wrx0 @ 1 = 1 @ q) =>
(fun (q0 : wry1 @ 1 = 1 @ y1)
(_ : (ehrnat_y1 [I] ulnat_y1) @ 1 = 1 @ q0) =>
(fun (q1 : wry0 @ 1 = 1 @ y0)
(_ : (ehrnat_y0 [I] ulnat_y0) @ 1 = 1 @ q1) =>
(fun (q2 : wlx1 @ 1 = 1 @ x1)
(_ : (ehlnat_x1 [I] urnat_x1) @ 1 = 1 @ q2) =>
(fun (q3 : wlx0 @ 1 = 1 @ q)
(_ : (ehlnat_x0 [I] s) @ 1 = 1 @ q3) =>
let EH_x_y :=
((rlucancel_inv (q3 [-] q1))^ @
(((concat_p1 (...))^ @ ((...) @ (...)^)) @
((concat_p1 (wry1 @ wlx1))^)^)) @
rlucancel_inv (q0 [-] q2) in
let EH_y_x :=
((rlucancel_inv (ulnat_y1 [-] urnat_x1))^ @
wlrnat_y_x) @ rlucancel_inv (s [-] ulnat_y0)
in
EH_x_y @ EH_y_x = 1 ) (ehlnat_x0 [I] s)
(rlucancel 1 )) (ehlnat_x1 [I] urnat_x1)
(rlucancel 1 )) (ehrnat_y0 [I] ulnat_y0)
(rlucancel 1 )) (ehrnat_y1 [I] ulnat_y1)
(rlucancel 1 )) wrx0 (rlucancel 1 )
revert x1 urnat_x1.X : Type a, b, c, d : X wlx0, wrx0 : a = b wlx1, wrx1 : c = d wly0, y0, wry0 : b = d wly1, y1, wry1 : a = c ulnat_y0 : wly0 @ 1 = 1 @ y0 ulnat_y1 : wly1 @ 1 = 1 @ y1 ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehrnat_y0 : wry0 @ 1 = 1 @ wly0 ehrnat_y1 : wry1 @ 1 = 1 @ wly1 wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0
forall (x1 : c = d) (urnat_x1 : wrx1 @ 1 = 1 @ x1),
(fun (q : a = b) (s : wrx0 @ 1 = 1 @ q) =>
(fun (q0 : wry1 @ 1 = 1 @ y1)
(_ : (ehrnat_y1 [I] ulnat_y1) @ 1 = 1 @ q0) =>
(fun (q1 : wry0 @ 1 = 1 @ y0)
(_ : (ehrnat_y0 [I] ulnat_y0) @ 1 = 1 @ q1) =>
(fun (q2 : wlx1 @ 1 = 1 @ x1)
(_ : (ehlnat_x1 [I] urnat_x1) @ 1 = 1 @ q2) =>
(fun (q3 : wlx0 @ 1 = 1 @ q)
(_ : (ehlnat_x0 [I] s) @ 1 = 1 @ q3) =>
let EH_x_y :=
((rlucancel_inv (q3 [-] q1))^ @
(((concat_p1 ...)^ @ (... @ ...^)) @
((concat_p1 (...))^)^)) @
rlucancel_inv (q0 [-] q2) in
let EH_y_x :=
((rlucancel_inv (ulnat_y1 [-] urnat_x1))^ @
wlrnat_y_x) @ rlucancel_inv (s [-] ulnat_y0)
in
EH_x_y @ EH_y_x = 1 ) (ehlnat_x0 [I] s)
(rlucancel 1 )) (ehlnat_x1 [I] urnat_x1)
(rlucancel 1 )) (ehrnat_y0 [I] ulnat_y0)
(rlucancel 1 )) (ehrnat_y1 [I] ulnat_y1)
(rlucancel 1 )) wrx0 (rlucancel 1 )
snapply equiv_path_ind_rlucancel. X : Type a, b, c, d : X wlx0, wrx0 : a = b wlx1, wrx1 : c = d wly0, y0, wry0 : b = d wly1, y1, wry1 : a = c ulnat_y0 : wly0 @ 1 = 1 @ y0 ulnat_y1 : wly1 @ 1 = 1 @ y1 ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehrnat_y0 : wry0 @ 1 = 1 @ wly0 ehrnat_y1 : wry1 @ 1 = 1 @ wly1 wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0
(fun (q : c = d) (s : wrx1 @ 1 = 1 @ q) =>
(fun (q0 : a = b) (s0 : wrx0 @ 1 = 1 @ q0) =>
(fun (q1 : wry1 @ 1 = 1 @ y1)
(_ : (ehrnat_y1 [I] ulnat_y1) @ 1 = 1 @ q1) =>
(fun (q2 : wry0 @ 1 = 1 @ y0)
(_ : (ehrnat_y0 [I] ulnat_y0) @ 1 = 1 @ q2) =>
(fun (q3 : wlx1 @ 1 = 1 @ q)
(_ : (ehlnat_x1 [I] s) @ 1 = 1 @ q3) =>
(fun (q4 : wlx0 @ 1 = 1 @ q0)
(_ : (ehlnat_x0 [I] s0) @ 1 = 1 @ q4) =>
let EH_x_y :=
((rlucancel_inv (q4 [-] q2))^ @
((...^ @ ...) @ ((...)^)^)) @
rlucancel_inv (q1 [-] q3) in
let EH_y_x :=
((rlucancel_inv (...))^ @ wlrnat_y_x) @
rlucancel_inv (s0 [-] ulnat_y0) in
EH_x_y @ EH_y_x = 1 ) (ehlnat_x0 [I] s0)
(rlucancel 1 )) (ehlnat_x1 [I] s) (rlucancel 1 ))
(ehrnat_y0 [I] ulnat_y0) (rlucancel 1 ))
(ehrnat_y1 [I] ulnat_y1) (rlucancel 1 )) wrx0
(rlucancel 1 )) wrx1 (rlucancel 1 )
revert y0 ulnat_y0.X : Type a, b, c, d : X wlx0, wrx0 : a = b wlx1, wrx1 : c = d wly0, wry0 : b = d wly1, y1, wry1 : a = c ulnat_y1 : wly1 @ 1 = 1 @ y1 ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehrnat_y0 : wry0 @ 1 = 1 @ wly0 ehrnat_y1 : wry1 @ 1 = 1 @ wly1 wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0
forall (y0 : b = d) (ulnat_y0 : wly0 @ 1 = 1 @ y0),
(fun (q : c = d) (s : wrx1 @ 1 = 1 @ q) =>
(fun (q0 : a = b) (s0 : wrx0 @ 1 = 1 @ q0) =>
(fun (q1 : wry1 @ 1 = 1 @ y1)
(_ : (ehrnat_y1 [I] ulnat_y1) @ 1 = 1 @ q1) =>
(fun (q2 : wry0 @ 1 = 1 @ y0)
(_ : (ehrnat_y0 [I] ulnat_y0) @ 1 = 1 @ q2) =>
(fun (q3 : wlx1 @ 1 = 1 @ q)
(_ : (ehlnat_x1 [I] s) @ 1 = 1 @ q3) =>
(fun (q4 : wlx0 @ 1 = 1 @ q0)
(_ : (ehlnat_x0 [I] s0) @ 1 = 1 @ q4) =>
let EH_x_y :=
((rlucancel_inv (...))^ @ ((...) @ (...^)^)) @
rlucancel_inv (q1 [-] q3) in
let EH_y_x :=
((rlucancel_inv ...)^ @ wlrnat_y_x) @
rlucancel_inv (s0 [-] ulnat_y0) in
EH_x_y @ EH_y_x = 1 ) (ehlnat_x0 [I] s0)
(rlucancel 1 )) (ehlnat_x1 [I] s) (rlucancel 1 ))
(ehrnat_y0 [I] ulnat_y0) (rlucancel 1 ))
(ehrnat_y1 [I] ulnat_y1) (rlucancel 1 )) wrx0
(rlucancel 1 )) wrx1 (rlucancel 1 )
snapply equiv_path_ind_rlucancel. X : Type a, b, c, d : X wlx0, wrx0 : a = b wlx1, wrx1 : c = d wly0, wry0 : b = d wly1, y1, wry1 : a = c ulnat_y1 : wly1 @ 1 = 1 @ y1 ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehrnat_y0 : wry0 @ 1 = 1 @ wly0 ehrnat_y1 : wry1 @ 1 = 1 @ wly1 wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0
(fun (q : b = d) (s : wly0 @ 1 = 1 @ q) =>
(fun (q0 : c = d) (s0 : wrx1 @ 1 = 1 @ q0) =>
(fun (q1 : a = b) (s1 : wrx0 @ 1 = 1 @ q1) =>
(fun (q2 : wry1 @ 1 = 1 @ y1)
(_ : (ehrnat_y1 [I] ulnat_y1) @ 1 = 1 @ q2) =>
(fun (q3 : wry0 @ 1 = 1 @ q)
(_ : (ehrnat_y0 [I] s) @ 1 = 1 @ q3) =>
(fun (q4 : wlx1 @ 1 = 1 @ q0)
(_ : (ehlnat_x1 [I] s0) @ 1 = 1 @ q4) =>
(fun (q5 : wlx0 @ 1 = 1 @ q1)
(_ : (ehlnat_x0 [I] s1) @ 1 = 1 @ q5) =>
let EH_x_y :=
((...)^ @ (...)) @ rlucancel_inv (q2 [-] q4)
in
let EH_y_x :=
(...^ @ wlrnat_y_x) @ rlucancel_inv (...) in
EH_x_y @ EH_y_x = 1 ) (ehlnat_x0 [I] s1)
(rlucancel 1 )) (ehlnat_x1 [I] s0)
(rlucancel 1 )) (ehrnat_y0 [I] s) (rlucancel 1 ))
(ehrnat_y1 [I] ulnat_y1) (rlucancel 1 )) wrx0
(rlucancel 1 )) wrx1 (rlucancel 1 )) wly0
(rlucancel 1 )
revert y1 ulnat_y1.X : Type a, b, c, d : X wlx0, wrx0 : a = b wlx1, wrx1 : c = d wly0, wry0 : b = d wly1, wry1 : a = c ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehrnat_y0 : wry0 @ 1 = 1 @ wly0 ehrnat_y1 : wry1 @ 1 = 1 @ wly1 wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0
forall (y1 : a = c) (ulnat_y1 : wly1 @ 1 = 1 @ y1),
(fun (q : b = d) (s : wly0 @ 1 = 1 @ q) =>
(fun (q0 : c = d) (s0 : wrx1 @ 1 = 1 @ q0) =>
(fun (q1 : a = b) (s1 : wrx0 @ 1 = 1 @ q1) =>
(fun (q2 : wry1 @ 1 = 1 @ y1)
(_ : (ehrnat_y1 [I] ulnat_y1) @ 1 = 1 @ q2) =>
(fun (q3 : wry0 @ 1 = 1 @ q)
(_ : (ehrnat_y0 [I] s) @ 1 = 1 @ q3) =>
(fun (q4 : wlx1 @ 1 = 1 @ q0)
(_ : (ehlnat_x1 [I] s0) @ 1 = 1 @ q4) =>
(fun (q5 : wlx0 @ 1 = 1 @ q1)
(_ : (...) @ 1 = 1 @ q5) =>
let EH_x_y :=
(...^ @ ...) @ rlucancel_inv (...) in
let EH_y_x := (...) @ rlucancel_inv ... in
EH_x_y @ EH_y_x = 1 ) (ehlnat_x0 [I] s1)
(rlucancel 1 )) (ehlnat_x1 [I] s0)
(rlucancel 1 )) (ehrnat_y0 [I] s) (rlucancel 1 ))
(ehrnat_y1 [I] ulnat_y1) (rlucancel 1 )) wrx0
(rlucancel 1 )) wrx1 (rlucancel 1 )) wly0
(rlucancel 1 )
snapply equiv_path_ind_rlucancel. X : Type a, b, c, d : X wlx0, wrx0 : a = b wlx1, wrx1 : c = d wly0, wry0 : b = d wly1, wry1 : a = c ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehrnat_y0 : wry0 @ 1 = 1 @ wly0 ehrnat_y1 : wry1 @ 1 = 1 @ wly1 wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0
(fun (q : a = c) (s : wly1 @ 1 = 1 @ q) =>
(fun (q0 : b = d) (s0 : wly0 @ 1 = 1 @ q0) =>
(fun (q1 : c = d) (s1 : wrx1 @ 1 = 1 @ q1) =>
(fun (q2 : a = b) (s2 : wrx0 @ 1 = 1 @ q2) =>
(fun (q3 : wry1 @ 1 = 1 @ q)
(_ : (ehrnat_y1 [I] s) @ 1 = 1 @ q3) =>
(fun (q4 : wry0 @ 1 = 1 @ q0)
(_ : (ehrnat_y0 [I] s0) @ 1 = 1 @ q4) =>
(fun (q5 : wlx1 @ 1 = 1 @ q1)
(_ : (ehlnat_x1 [I] s1) @ 1 = 1 @ q5) =>
(fun (q6 : ... = ...) (_ : ... = ...) =>
let EH_x_y := ... @ ... in
let EH_y_x := ... in ... = 1 )
(ehlnat_x0 [I] s2) (rlucancel 1 ))
(ehlnat_x1 [I] s1) (rlucancel 1 ))
(ehrnat_y0 [I] s0) (rlucancel 1 ))
(ehrnat_y1 [I] s) (rlucancel 1 )) wrx0
(rlucancel 1 )) wrx1 (rlucancel 1 )) wly0
(rlucancel 1 )) wly1 (rlucancel 1 )
revert wlrnat_y_x.X : Type a, b, c, d : X wlx0, wrx0 : a = b wlx1, wrx1 : c = d wly0, wry0 : b = d wly1, wry1 : a = c ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehrnat_y0 : wry0 @ 1 = 1 @ wly0 ehrnat_y1 : wry1 @ 1 = 1 @ wly1
forall wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0,
(fun (q : a = c) (s : wly1 @ 1 = 1 @ q) =>
(fun (q0 : b = d) (s0 : wly0 @ 1 = 1 @ q0) =>
(fun (q1 : c = d) (s1 : wrx1 @ 1 = 1 @ q1) =>
(fun (q2 : a = b) (s2 : wrx0 @ 1 = 1 @ q2) =>
(fun (q3 : wry1 @ 1 = 1 @ q)
(_ : (ehrnat_y1 [I] s) @ 1 = 1 @ q3) =>
(fun (q4 : wry0 @ 1 = 1 @ q0)
(_ : (ehrnat_y0 [I] s0) @ 1 = 1 @ q4) =>
(fun (q5 : wlx1 @ 1 = 1 @ q1)
(_ : (...) @ 1 = 1 @ q5) =>
(fun (q6 : ...) (_ : ...) =>
let EH_x_y := ... in ... ...)
(ehlnat_x0 [I] s2) (rlucancel 1 ))
(ehlnat_x1 [I] s1) (rlucancel 1 ))
(ehrnat_y0 [I] s0) (rlucancel 1 ))
(ehrnat_y1 [I] s) (rlucancel 1 )) wrx0
(rlucancel 1 )) wrx1 (rlucancel 1 )) wly0
(rlucancel 1 )) wly1 (rlucancel 1 )
revert wrx0 ehlnat_x0.X : Type a, b, c, d : X wlx0 : a = b wlx1, wrx1 : c = d wly0, wry0 : b = d wly1, wry1 : a = c ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehrnat_y0 : wry0 @ 1 = 1 @ wly0 ehrnat_y1 : wry1 @ 1 = 1 @ wly1
forall (wrx0 : a = b)
(ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0)
(wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0),
(fun (q : a = c) (s : wly1 @ 1 = 1 @ q) =>
(fun (q0 : b = d) (s0 : wly0 @ 1 = 1 @ q0) =>
(fun (q1 : c = d) (s1 : wrx1 @ 1 = 1 @ q1) =>
(fun (q2 : a = b) (s2 : wrx0 @ 1 = 1 @ q2) =>
(fun (q3 : wry1 @ 1 = 1 @ q)
(_ : (ehrnat_y1 [I] s) @ 1 = 1 @ q3) =>
(fun (q4 : wry0 @ 1 = 1 @ q0)
(_ : (ehrnat_y0 [I] s0) @ 1 = 1 @ q4) =>
(fun (q5 : wlx1 @ 1 = 1 @ q1)
(_ : (...) @ 1 = 1 @ q5) =>
(fun (q6 : ...) (_ : ...) =>
let EH_x_y := ... in ... ...)
(ehlnat_x0 [I] s2) (rlucancel 1 ))
(ehlnat_x1 [I] s1) (rlucancel 1 ))
(ehrnat_y0 [I] s0) (rlucancel 1 ))
(ehrnat_y1 [I] s) (rlucancel 1 )) wrx0
(rlucancel 1 )) wrx1 (rlucancel 1 )) wly0
(rlucancel 1 )) wly1 (rlucancel 1 )
snapply equiv_path_ind_rlucancel. X : Type a, b, c, d : X wlx0 : a = b wlx1, wrx1 : c = d wly0, wry0 : b = d wly1, wry1 : a = c ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehrnat_y0 : wry0 @ 1 = 1 @ wly0 ehrnat_y1 : wry1 @ 1 = 1 @ wly1
(fun (q : a = b) (s : wlx0 @ 1 = 1 @ q) =>
forall wlrnat_y_x : wly1 @ wrx1 = q @ wly0,
(fun (q0 : a = c) (s0 : wly1 @ 1 = 1 @ q0) =>
(fun (q1 : b = d) (s1 : wly0 @ 1 = 1 @ q1) =>
(fun (q2 : c = d) (s2 : wrx1 @ 1 = 1 @ q2) =>
(fun (q3 : a = b) (s3 : q @ 1 = 1 @ q3) =>
(fun (q4 : wry1 @ 1 = 1 @ q0)
(_ : (ehrnat_y1 [I] s0) @ 1 = 1 @ q4) =>
(fun (q5 : wry0 @ 1 = 1 @ q1)
(_ : (...) @ 1 = 1 @ q5) =>
(fun (q6 : ...) (_ : ...) => (...) (...) (...))
(ehlnat_x1 [I] s2) (rlucancel 1 ))
(ehrnat_y0 [I] s1) (rlucancel 1 ))
(ehrnat_y1 [I] s0) (rlucancel 1 )) q
(rlucancel 1 )) wrx1 (rlucancel 1 )) wly0
(rlucancel 1 )) wly1 (rlucancel 1 )) wlx0
(rlucancel 1 )
revert wrx1 ehlnat_x1.X : Type a, b, c, d : X wlx0 : a = b wlx1 : c = d wly0, wry0 : b = d wly1, wry1 : a = c ehrnat_y0 : wry0 @ 1 = 1 @ wly0 ehrnat_y1 : wry1 @ 1 = 1 @ wly1
forall (wrx1 : c = d)
(ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1),
(fun (q : a = b) (s : wlx0 @ 1 = 1 @ q) =>
forall wlrnat_y_x : wly1 @ wrx1 = q @ wly0,
(fun (q0 : a = c) (s0 : wly1 @ 1 = 1 @ q0) =>
(fun (q1 : b = d) (s1 : wly0 @ 1 = 1 @ q1) =>
(fun (q2 : c = d) (s2 : wrx1 @ 1 = 1 @ q2) =>
(fun (q3 : a = b) (s3 : q @ 1 = 1 @ q3) =>
(fun (q4 : wry1 @ 1 = 1 @ q0)
(_ : (ehrnat_y1 [I] s0) @ 1 = 1 @ q4) =>
(fun (q5 : wry0 @ 1 = 1 @ q1)
(_ : ... @ 1 = 1 @ q5) =>
(fun ... ... => ... ... ...) (ehlnat_x1 [I] s2)
(rlucancel 1 )) (ehrnat_y0 [I] s1)
(rlucancel 1 )) (ehrnat_y1 [I] s0)
(rlucancel 1 )) q (rlucancel 1 )) wrx1
(rlucancel 1 )) wly0 (rlucancel 1 )) wly1
(rlucancel 1 )) wlx0 (rlucancel 1 )
snapply equiv_path_ind_rlucancel. X : Type a, b, c, d : X wlx0 : a = b wlx1 : c = d wly0, wry0 : b = d wly1, wry1 : a = c ehrnat_y0 : wry0 @ 1 = 1 @ wly0 ehrnat_y1 : wry1 @ 1 = 1 @ wly1
(fun (q : c = d) (s : wlx1 @ 1 = 1 @ q) =>
(fun (q0 : a = b) (s0 : wlx0 @ 1 = 1 @ q0) =>
forall wlrnat_y_x : wly1 @ q = q0 @ wly0,
(fun (q1 : a = c) (s1 : wly1 @ 1 = 1 @ q1) =>
(fun (q2 : b = d) (s2 : wly0 @ 1 = 1 @ q2) =>
(fun (q3 : c = d) (s3 : q @ 1 = 1 @ q3) =>
(fun (q4 : a = b) (s4 : q0 @ 1 = 1 @ q4) =>
(fun (q5 : wry1 @ 1 = 1 @ q1)
(_ : (...) @ 1 = 1 @ q5) =>
(fun (q6 : ...) (_ : ...) => (...) (...) (...))
(ehrnat_y0 [I] s2) (rlucancel 1 ))
(ehrnat_y1 [I] s1) (rlucancel 1 )) q0
(rlucancel 1 )) q (rlucancel 1 )) wly0
(rlucancel 1 )) wly1 (rlucancel 1 )) wlx0
(rlucancel 1 )) wlx1 (rlucancel 1 )
revert wly0 ehrnat_y0.X : Type a, b, c, d : X wlx0 : a = b wlx1 : c = d wry0 : b = d wly1, wry1 : a = c ehrnat_y1 : wry1 @ 1 = 1 @ wly1
forall (wly0 : b = d)
(ehrnat_y0 : wry0 @ 1 = 1 @ wly0),
(fun (q : c = d) (s : wlx1 @ 1 = 1 @ q) =>
(fun (q0 : a = b) (s0 : wlx0 @ 1 = 1 @ q0) =>
forall wlrnat_y_x : wly1 @ q = q0 @ wly0,
(fun (q1 : a = c) (s1 : wly1 @ 1 = 1 @ q1) =>
(fun (q2 : b = d) (s2 : wly0 @ 1 = 1 @ q2) =>
(fun (q3 : c = d) (s3 : q @ 1 = 1 @ q3) =>
(fun (q4 : a = b) (s4 : q0 @ 1 = 1 @ q4) =>
(fun (q5 : wry1 @ 1 = 1 @ q1)
(_ : ... @ 1 = 1 @ q5) =>
(fun ... ... => ... ... ...) (ehrnat_y0 [I] s2)
(rlucancel 1 )) (ehrnat_y1 [I] s1)
(rlucancel 1 )) q0 (rlucancel 1 )) q
(rlucancel 1 )) wly0 (rlucancel 1 )) wly1
(rlucancel 1 )) wlx0 (rlucancel 1 )) wlx1
(rlucancel 1 )
snapply equiv_path_ind_rlucancel. X : Type a, b, c, d : X wlx0 : a = b wlx1 : c = d wry0 : b = d wly1, wry1 : a = c ehrnat_y1 : wry1 @ 1 = 1 @ wly1
(fun (q : b = d) (s : wry0 @ 1 = 1 @ q) =>
(fun (q0 : c = d) (s0 : wlx1 @ 1 = 1 @ q0) =>
(fun (q1 : a = b) (s1 : wlx0 @ 1 = 1 @ q1) =>
forall wlrnat_y_x : wly1 @ q0 = q1 @ q,
(fun (q2 : a = c) (s2 : wly1 @ 1 = 1 @ q2) =>
(fun (q3 : b = d) (s3 : q @ 1 = 1 @ q3) =>
(fun (q4 : c = d) (s4 : q0 @ 1 = 1 @ q4) =>
(fun (q5 : a = b) (s5 : q1 @ 1 = 1 @ q5) =>
(fun (q6 : ...) (_ : ...) => (...) (...) (...))
(ehrnat_y1 [I] s2) (rlucancel 1 )) q1
(rlucancel 1 )) q0 (rlucancel 1 )) q
(rlucancel 1 )) wly1 (rlucancel 1 )) wlx0
(rlucancel 1 )) wlx1 (rlucancel 1 )) wry0
(rlucancel 1 )
revert wly1 ehrnat_y1.X : Type a, b, c, d : X wlx0 : a = b wlx1 : c = d wry0 : b = d wry1 : a = c
forall (wly1 : a = c)
(ehrnat_y1 : wry1 @ 1 = 1 @ wly1),
(fun (q : b = d) (s : wry0 @ 1 = 1 @ q) =>
(fun (q0 : c = d) (s0 : wlx1 @ 1 = 1 @ q0) =>
(fun (q1 : a = b) (s1 : wlx0 @ 1 = 1 @ q1) =>
forall wlrnat_y_x : wly1 @ q0 = q1 @ q,
(fun (q2 : a = c) (s2 : wly1 @ 1 = 1 @ q2) =>
(fun (q3 : b = d) (s3 : q @ 1 = 1 @ q3) =>
(fun (q4 : c = d) (s4 : q0 @ 1 = 1 @ q4) =>
(fun (q5 : a = b) (s5 : q1 @ 1 = 1 @ q5) =>
(fun ... ... => ... ... ...) (ehrnat_y1 [I] s2)
(rlucancel 1 )) q1 (rlucancel 1 )) q0
(rlucancel 1 )) q (rlucancel 1 )) wly1
(rlucancel 1 )) wlx0 (rlucancel 1 )) wlx1
(rlucancel 1 )) wry0 (rlucancel 1 )
snapply equiv_path_ind_rlucancel. X : Type a, b, c, d : X wlx0 : a = b wlx1 : c = d wry0 : b = d wry1 : a = c
(fun (q : a = c) (s : wry1 @ 1 = 1 @ q) =>
(fun (q0 : b = d) (s0 : wry0 @ 1 = 1 @ q0) =>
(fun (q1 : c = d) (s1 : wlx1 @ 1 = 1 @ q1) =>
(fun (q2 : a = b) (s2 : wlx0 @ 1 = 1 @ q2) =>
forall wlrnat_y_x : q @ q1 = q2 @ q0,
(fun (q3 : a = c) (s3 : q @ 1 = 1 @ q3) =>
(fun (q4 : b = d) (s4 : q0 @ 1 = 1 @ q4) =>
(fun (q5 : c = d) (s5 : q1 @ 1 = 1 @ q5) =>
(fun (q6 : ...) (s6 : ...) => (...) (...) (...))
q2 (rlucancel 1 )) q1 (rlucancel 1 )) q0
(rlucancel 1 )) q (rlucancel 1 )) wlx0
(rlucancel 1 )) wlx1 (rlucancel 1 )) wry0
(rlucancel 1 )) wry1 (rlucancel 1 )
destruct wry0, wry1, wlx1.X : Type a : X wlx0, wlx1, wry0, wry1 : a = a
forall wlrnat_y_x : 1 @ 1 = wlx0 @ 1 ,
let EH_x_y :=
((rlucancel_inv
((rlucancel 1 [I] rlucancel 1 )
[-] (rlucancel 1 [I] rlucancel 1 )))^ @
(((concat_p1 (wlx0 @ 1 ))^ @
(((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 wlrnat_y_x^) @
(rlucancel 1 [-] rlucancel 1 )^)) @
((concat_p1 (1 @ 1 ))^)^)) @
rlucancel_inv
((rlucancel 1 [I] rlucancel 1 )
[-] (rlucancel 1 [I] rlucancel 1 )) in
let EH_y_x :=
((rlucancel_inv (rlucancel 1 [-] rlucancel 1 ))^ @
wlrnat_y_x) @
rlucancel_inv (rlucancel 1 [-] rlucancel 1 ) in
EH_x_y @ EH_y_x = 1
clear wry0 wry1 wlx1.X : Type a : X wlx0 : a = a
forall wlrnat_y_x : 1 @ 1 = wlx0 @ 1 ,
let EH_x_y :=
((rlucancel_inv
((rlucancel 1 [I] rlucancel 1 )
[-] (rlucancel 1 [I] rlucancel 1 )))^ @
(((concat_p1 (wlx0 @ 1 ))^ @
(((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 wlrnat_y_x^) @
(rlucancel 1 [-] rlucancel 1 )^)) @
((concat_p1 (1 @ 1 ))^)^)) @
rlucancel_inv
((rlucancel 1 [I] rlucancel 1 )
[-] (rlucancel 1 [I] rlucancel 1 )) in
let EH_y_x :=
((rlucancel_inv (rlucancel 1 [-] rlucancel 1 ))^ @
wlrnat_y_x) @
rlucancel_inv (rlucancel 1 [-] rlucancel 1 ) in
EH_x_y @ EH_y_x = 1
revert wlx0.X : Type a : X
forall (wlx0 : a = a) (wlrnat_y_x : 1 @ 1 = wlx0 @ 1 ),
let EH_x_y :=
((rlucancel_inv
((rlucancel 1 [I] rlucancel 1 )
[-] (rlucancel 1 [I] rlucancel 1 )))^ @
(((concat_p1 (wlx0 @ 1 ))^ @
(((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 wlrnat_y_x^) @
(rlucancel 1 [-] rlucancel 1 )^)) @
((concat_p1 (1 @ 1 ))^)^)) @
rlucancel_inv
((rlucancel 1 [I] rlucancel 1 )
[-] (rlucancel 1 [I] rlucancel 1 )) in
let EH_y_x :=
((rlucancel_inv (rlucancel 1 [-] rlucancel 1 ))^ @
wlrnat_y_x) @
rlucancel_inv (rlucancel 1 [-] rlucancel 1 ) in
EH_x_y @ EH_y_x = 1
snapply equiv_path_ind_lrucancel. X : Type a : X
(fun (q : a = a) (s : 1 @ 1 = q @ 1 ) =>
let EH_x_y :=
((rlucancel_inv
((rlucancel 1 [I] rlucancel 1 )
[-] (rlucancel 1 [I] rlucancel 1 )))^ @
(((concat_p1 (q @ 1 ))^ @
(((rlucancel 1 [-] rlucancel 1 ) @ whiskerL 1 s^) @
(rlucancel 1 [-] rlucancel 1 )^)) @
((concat_p1 (1 @ 1 ))^)^)) @
rlucancel_inv
((rlucancel 1 [I] rlucancel 1 )
[-] (rlucancel 1 [I] rlucancel 1 )) in
let EH_y_x :=
((rlucancel_inv (rlucancel 1 [-] rlucancel 1 ))^ @ s) @
rlucancel_inv (rlucancel 1 [-] rlucancel 1 ) in
EH_x_y @ EH_y_x = 1 ) 1 (lrucancel 1 )
reflexivity .
Defined .
End eh_V .
Theorem eh_V {X } {a : X} (p q : idpath (idpath a) = idpath (idpath a)) :
eh p q @ eh q p = 1 .X : Type a : X p, q : 1 = 1
eh p q @ eh q p = 1
Proof .X : Type a : X p, q : 1 = 1
eh p q @ eh q p = 1
napply eh_V_gen. X : Type a : X p, q : 1 = 1
(?Goal [I] urnat p) @ 1 = 1 @ ulnat p
- X : Type a : X p, q : 1 = 1
(?Goal [I] urnat p) @ 1 = 1 @ ulnat p
exact (ehlnat_1p p).
- X : Type a : X p, q : 1 = 1
(?Goal [I] urnat p) @ 1 = 1 @ ulnat p
exact (ehlnat_1p p).
- X : Type a : X p, q : 1 = 1
(?Goal [I] ulnat q) @ 1 = 1 @ urnat q
exact (ehrnat_p1 q).
- X : Type a : X p, q : 1 = 1
(?Goal [I] ulnat q) @ 1 = 1 @ urnat q
exact (ehrnat_p1 q).
- X : Type a : X p, q : 1 = 1
whiskerR (wlrnat p q) 1 @ (ehrnat q 1 [-] ehlnat 1 p) =
(ehlnat 1 p [-] ehrnat q 1 ) @ whiskerL 1 (wlrnat q p)^
exact (wlrnat_V p q).
Defined .
(** Given [ehrnat_p1 y] and [ehrnat_p1 z], we can explicitly construct [ehrnat_p1 (y @ z)]. *)
Section Ehrnat_p1_pp .
Context {X : Type }.
(** 0-paths *)
Context {a0 a1 a2 : X}.
Context {b0 b1 b2 : X}.
Context {c0 c1 c2 : X}.
(** 1-paths *)
Context {wry : a0 = b0}.
Context {wrz : b0 = c0}.
Context {wly : a1 = b1}.
Context {wlz : b1 = c1}.
Context {y : a2 = b2}.
Context {z : b2 = c2}.
Context {wryz : a0 = c0}.
Context {wlyz : a1 = c1}.
Context {a01 : a0 = a1}.
Context {a12 : a1 = a2}.
Context {b01 : b0 = b1}.
Context {b12 : b1 = b2}.
Context {c01 : c0 = c1}.
Context {c12 : c1 = c2}.
Context {a02 : a0 = a2}.
Context {c02 : c0 = c2}.
(** 2-paths *)
Context {ehrnat_y : wry @ b01 = a01 @ wly}.
Context {ehrnat_z : wrz @ c01 = b01 @ wlz}.
Context {ehrnat_yz : wryz @ c01 = a01 @ wlyz}.
Context {ulnat_y : wly @ b12 = a12 @ y}.
Context {ulnat_z : wlz @ c12 = b12 @ z}.
Context {ulnat_yz : wlyz @ c12 = a12 @ (y @ z)}.
Context {urnat_y : wry @ (b01 @ b12) = a02 @ y}.
Context {urnat_z : wrz @ c02 = (b01 @ b12) @ z}.
Context {urnat_yz : wryz @ c02 = a02 @ (y @ z)}.
Context {wrpp_yz : wry @ wrz = wryz}.
Context {wlpp_yz : wly @ wlz = wlyz}.
Context (H_a02 : a01 @ a12 = a02).
Context (H_c02 : c01 @ c12 = c02).
(*& 3-paths *)
Hypothesis H_ehrnat_yz :
(ehrnat_y [-] ehrnat_z) @ whiskerL _ wlpp_yz =
whiskerR wrpp_yz _ @ ehrnat_yz.
Hypothesis H_ulnat_yz :
(ulnat_y [-] ulnat_z) = whiskerR wlpp_yz _ @ ulnat_yz.
Hypothesis H_urnat_yz :
(urnat_y [-] urnat_z) = whiskerR wrpp_yz _ @ urnat_yz.
Variable ehrnat_p1_y :
(ehrnat_y [I] ulnat_y) @ whiskerR H_a02 _ = 1 @ urnat_y.
Variable ehrnat_p1_z :
(ehrnat_z [I] ulnat_z) @ 1 = whiskerL _ H_c02 @ urnat_z.
(** The composite iso *)
Definition Ehrnat_p1_pp :
(ehrnat_yz [I] ulnat_yz) @ whiskerR H_a02 _ =
whiskerL _ H_c02 @ urnat_yz.X : Type a0, a1, a2, b0, b1, b2, c0, c1, c2 : X wry : a0 = b0 wrz : b0 = c0 wly : a1 = b1 wlz : b1 = c1 y : a2 = b2 z : b2 = c2 wryz : a0 = c0 wlyz : a1 = c1 a01 : a0 = a1 a12 : a1 = a2 b01 : b0 = b1 b12 : b1 = b2 c01 : c0 = c1 c12 : c1 = c2 a02 : a0 = a2 c02 : c0 = c2 ehrnat_y : wry @ b01 = a01 @ wly ehrnat_z : wrz @ c01 = b01 @ wlz ehrnat_yz : wryz @ c01 = a01 @ wlyz ulnat_y : wly @ b12 = a12 @ y ulnat_z : wlz @ c12 = b12 @ z ulnat_yz : wlyz @ c12 = a12 @ (y @ z) urnat_y : wry @ (b01 @ b12) = a02 @ y urnat_z : wrz @ c02 = (b01 @ b12) @ z urnat_yz : wryz @ c02 = a02 @ (y @ z) wrpp_yz : wry @ wrz = wryz wlpp_yz : wly @ wlz = wlyz H_a02 : a01 @ a12 = a02 H_c02 : c01 @ c12 = c02 H_ehrnat_yz : (ehrnat_y [-] ehrnat_z) @
whiskerL a01 wlpp_yz =
whiskerR wrpp_yz c01 @ ehrnat_yz H_ulnat_yz : ulnat_y [-] ulnat_z =
whiskerR wlpp_yz c12 @ ulnat_yz H_urnat_yz : urnat_y [-] urnat_z =
whiskerR wrpp_yz c02 @ urnat_yz ehrnat_p1_y : (ehrnat_y [I] ulnat_y) @
whiskerR H_a02 y = 1 @ urnat_y ehrnat_p1_z : (ehrnat_z [I] ulnat_z) @ 1 =
whiskerL wrz H_c02 @ urnat_z
(ehrnat_yz [I] ulnat_yz) @ whiskerR H_a02 (y @ z) =
whiskerL wryz H_c02 @ urnat_yz
Proof .X : Type a0, a1, a2, b0, b1, b2, c0, c1, c2 : X wry : a0 = b0 wrz : b0 = c0 wly : a1 = b1 wlz : b1 = c1 y : a2 = b2 z : b2 = c2 wryz : a0 = c0 wlyz : a1 = c1 a01 : a0 = a1 a12 : a1 = a2 b01 : b0 = b1 b12 : b1 = b2 c01 : c0 = c1 c12 : c1 = c2 a02 : a0 = a2 c02 : c0 = c2 ehrnat_y : wry @ b01 = a01 @ wly ehrnat_z : wrz @ c01 = b01 @ wlz ehrnat_yz : wryz @ c01 = a01 @ wlyz ulnat_y : wly @ b12 = a12 @ y ulnat_z : wlz @ c12 = b12 @ z ulnat_yz : wlyz @ c12 = a12 @ (y @ z) urnat_y : wry @ (b01 @ b12) = a02 @ y urnat_z : wrz @ c02 = (b01 @ b12) @ z urnat_yz : wryz @ c02 = a02 @ (y @ z) wrpp_yz : wry @ wrz = wryz wlpp_yz : wly @ wlz = wlyz H_a02 : a01 @ a12 = a02 H_c02 : c01 @ c12 = c02 H_ehrnat_yz : (ehrnat_y [-] ehrnat_z) @
whiskerL a01 wlpp_yz =
whiskerR wrpp_yz c01 @ ehrnat_yz H_ulnat_yz : ulnat_y [-] ulnat_z =
whiskerR wlpp_yz c12 @ ulnat_yz H_urnat_yz : urnat_y [-] urnat_z =
whiskerR wrpp_yz c02 @ urnat_yz ehrnat_p1_y : (ehrnat_y [I] ulnat_y) @
whiskerR H_a02 y = 1 @ urnat_y ehrnat_p1_z : (ehrnat_z [I] ulnat_z) @ 1 =
whiskerL wrz H_c02 @ urnat_z
(ehrnat_yz [I] ulnat_yz) @ whiskerR H_a02 (y @ z) =
whiskerL wryz H_c02 @ urnat_yz
apply moveR_Vp in H_urnat_yz, H_ulnat_yz, H_ehrnat_yz.X : Type a0, a1, a2, b0, b1, b2, c0, c1, c2 : X wry : a0 = b0 wrz : b0 = c0 wly : a1 = b1 wlz : b1 = c1 y : a2 = b2 z : b2 = c2 wryz : a0 = c0 wlyz : a1 = c1 a01 : a0 = a1 a12 : a1 = a2 b01 : b0 = b1 b12 : b1 = b2 c01 : c0 = c1 c12 : c1 = c2 a02 : a0 = a2 c02 : c0 = c2 ehrnat_y : wry @ b01 = a01 @ wly ehrnat_z : wrz @ c01 = b01 @ wlz ehrnat_yz : wryz @ c01 = a01 @ wlyz ulnat_y : wly @ b12 = a12 @ y ulnat_z : wlz @ c12 = b12 @ z ulnat_yz : wlyz @ c12 = a12 @ (y @ z) urnat_y : wry @ (b01 @ b12) = a02 @ y urnat_z : wrz @ c02 = (b01 @ b12) @ z urnat_yz : wryz @ c02 = a02 @ (y @ z) wrpp_yz : wry @ wrz = wryz wlpp_yz : wly @ wlz = wlyz H_a02 : a01 @ a12 = a02 H_c02 : c01 @ c12 = c02 H_ehrnat_yz : (whiskerR wrpp_yz c01)^ @
((ehrnat_y [-] ehrnat_z) @
whiskerL a01 wlpp_yz) = ehrnat_yz H_ulnat_yz : (whiskerR wlpp_yz c12)^ @
(ulnat_y [-] ulnat_z) = ulnat_yz H_urnat_yz : (whiskerR wrpp_yz c02)^ @
(urnat_y [-] urnat_z) = urnat_yz ehrnat_p1_y : (ehrnat_y [I] ulnat_y) @
whiskerR H_a02 y = 1 @ urnat_y ehrnat_p1_z : (ehrnat_z [I] ulnat_z) @ 1 =
whiskerL wrz H_c02 @ urnat_z
(ehrnat_yz [I] ulnat_yz) @ whiskerR H_a02 (y @ z) =
whiskerL wryz H_c02 @ urnat_yz
destruct H_urnat_yz, H_ulnat_yz, H_ehrnat_yz.X : Type a0, a1, a2, b0, b1, b2, c0, c1, c2 : X wry : a0 = b0 wrz : b0 = c0 wly : a1 = b1 wlz : b1 = c1 y : a2 = b2 z : b2 = c2 wryz : a0 = c0 wlyz : a1 = c1 a01 : a0 = a1 a12 : a1 = a2 b01 : b0 = b1 b12 : b1 = b2 c01 : c0 = c1 c12 : c1 = c2 a02 : a0 = a2 c02 : c0 = c2 ehrnat_y : wry @ b01 = a01 @ wly ehrnat_z : wrz @ c01 = b01 @ wlz ulnat_y : wly @ b12 = a12 @ y ulnat_z : wlz @ c12 = b12 @ z urnat_y : wry @ (b01 @ b12) = a02 @ y urnat_z : wrz @ c02 = (b01 @ b12) @ z wrpp_yz : wry @ wrz = wryz wlpp_yz : wly @ wlz = wlyz H_a02 : a01 @ a12 = a02 H_c02 : c01 @ c12 = c02 H_ehrnat_yz : (whiskerR wrpp_yz c01)^ @
((ehrnat_y [-] ehrnat_z) @
whiskerL a01 wlpp_yz) =
(whiskerR wrpp_yz c01)^ @
((ehrnat_y [-] ehrnat_z) @
whiskerL a01 wlpp_yz) H_ulnat_yz : (whiskerR wlpp_yz c12)^ @
(ulnat_y [-] ulnat_z) =
(whiskerR wlpp_yz c12)^ @
(ulnat_y [-] ulnat_z) H_urnat_yz : (whiskerR wrpp_yz c02)^ @
(urnat_y [-] urnat_z) =
(whiskerR wrpp_yz c02)^ @
(urnat_y [-] urnat_z) ehrnat_p1_y : (ehrnat_y [I] ulnat_y) @
whiskerR H_a02 y = 1 @ urnat_y ehrnat_p1_z : (ehrnat_z [I] ulnat_z) @ 1 =
whiskerL wrz H_c02 @ urnat_z
((whiskerR wrpp_yz c01)^ @
((ehrnat_y [-] ehrnat_z) @ whiskerL a01 wlpp_yz)
[I] (whiskerR wlpp_yz c12)^ @ (ulnat_y [-] ulnat_z)) @
whiskerR H_a02 (y @ z) =
whiskerL wryz H_c02 @
((whiskerR wrpp_yz c02)^ @ (urnat_y [-] urnat_z))
clear H_urnat_yz H_ulnat_yz H_ehrnat_yz.X : Type a0, a1, a2, b0, b1, b2, c0, c1, c2 : X wry : a0 = b0 wrz : b0 = c0 wly : a1 = b1 wlz : b1 = c1 y : a2 = b2 z : b2 = c2 wryz : a0 = c0 wlyz : a1 = c1 a01 : a0 = a1 a12 : a1 = a2 b01 : b0 = b1 b12 : b1 = b2 c01 : c0 = c1 c12 : c1 = c2 a02 : a0 = a2 c02 : c0 = c2 ehrnat_y : wry @ b01 = a01 @ wly ehrnat_z : wrz @ c01 = b01 @ wlz ulnat_y : wly @ b12 = a12 @ y ulnat_z : wlz @ c12 = b12 @ z urnat_y : wry @ (b01 @ b12) = a02 @ y urnat_z : wrz @ c02 = (b01 @ b12) @ z wrpp_yz : wry @ wrz = wryz wlpp_yz : wly @ wlz = wlyz H_a02 : a01 @ a12 = a02 H_c02 : c01 @ c12 = c02 ehrnat_p1_y : (ehrnat_y [I] ulnat_y) @
whiskerR H_a02 y = 1 @ urnat_y ehrnat_p1_z : (ehrnat_z [I] ulnat_z) @ 1 =
whiskerL wrz H_c02 @ urnat_z
((whiskerR wrpp_yz c01)^ @
((ehrnat_y [-] ehrnat_z) @ whiskerL a01 wlpp_yz)
[I] (whiskerR wlpp_yz c12)^ @ (ulnat_y [-] ulnat_z)) @
whiskerR H_a02 (y @ z) =
whiskerL wryz H_c02 @
((whiskerR wrpp_yz c02)^ @ (urnat_y [-] urnat_z))
apply moveR_Vp in ehrnat_p1_y, ehrnat_p1_z.X : Type a0, a1, a2, b0, b1, b2, c0, c1, c2 : X wry : a0 = b0 wrz : b0 = c0 wly : a1 = b1 wlz : b1 = c1 y : a2 = b2 z : b2 = c2 wryz : a0 = c0 wlyz : a1 = c1 a01 : a0 = a1 a12 : a1 = a2 b01 : b0 = b1 b12 : b1 = b2 c01 : c0 = c1 c12 : c1 = c2 a02 : a0 = a2 c02 : c0 = c2 ehrnat_y : wry @ b01 = a01 @ wly ehrnat_z : wrz @ c01 = b01 @ wlz ulnat_y : wly @ b12 = a12 @ y ulnat_z : wlz @ c12 = b12 @ z urnat_y : wry @ (b01 @ b12) = a02 @ y urnat_z : wrz @ c02 = (b01 @ b12) @ z wrpp_yz : wry @ wrz = wryz wlpp_yz : wly @ wlz = wlyz H_a02 : a01 @ a12 = a02 H_c02 : c01 @ c12 = c02 ehrnat_p1_y : 1 ^ @
((ehrnat_y [I] ulnat_y) @
whiskerR H_a02 y) = urnat_yehrnat_p1_z : (whiskerL wrz H_c02)^ @
((ehrnat_z [I] ulnat_z) @ 1 ) = urnat_z
((whiskerR wrpp_yz c01)^ @
((ehrnat_y [-] ehrnat_z) @ whiskerL a01 wlpp_yz)
[I] (whiskerR wlpp_yz c12)^ @ (ulnat_y [-] ulnat_z)) @
whiskerR H_a02 (y @ z) =
whiskerL wryz H_c02 @
((whiskerR wrpp_yz c02)^ @ (urnat_y [-] urnat_z))
destruct ehrnat_p1_y, ehrnat_p1_z.X : Type a0, a1, a2, b0, b1, b2, c0, c1, c2 : X wry : a0 = b0 wrz : b0 = c0 wly : a1 = b1 wlz : b1 = c1 y : a2 = b2 z : b2 = c2 wryz : a0 = c0 wlyz : a1 = c1 a01 : a0 = a1 a12 : a1 = a2 b01 : b0 = b1 b12 : b1 = b2 c01 : c0 = c1 c12 : c1 = c2 a02 : a0 = a2 c02 : c0 = c2 ehrnat_y : wry @ b01 = a01 @ wly ehrnat_z : wrz @ c01 = b01 @ wlz ulnat_y : wly @ b12 = a12 @ y ulnat_z : wlz @ c12 = b12 @ z wrpp_yz : wry @ wrz = wryz wlpp_yz : wly @ wlz = wlyz H_a02 : a01 @ a12 = a02 H_c02 : c01 @ c12 = c02 ehrnat_p1_y : 1 ^ @
((ehrnat_y [I] ulnat_y) @
whiskerR H_a02 y) =
1 ^ @
((ehrnat_y [I] ulnat_y) @
whiskerR H_a02 y)ehrnat_p1_z : (whiskerL wrz H_c02)^ @
((ehrnat_z [I] ulnat_z) @ 1 ) =
(whiskerL wrz H_c02)^ @
((ehrnat_z [I] ulnat_z) @ 1 )
((whiskerR wrpp_yz c01)^ @
((ehrnat_y [-] ehrnat_z) @ whiskerL a01 wlpp_yz)
[I] (whiskerR wlpp_yz c12)^ @ (ulnat_y [-] ulnat_z)) @
whiskerR H_a02 (y @ z) =
whiskerL wryz H_c02 @
((whiskerR wrpp_yz c02)^ @
(1 ^ @ ((ehrnat_y [I] ulnat_y) @ whiskerR H_a02 y)
[-] (whiskerL wrz H_c02)^ @
((ehrnat_z [I] ulnat_z) @ 1 )))
clear ehrnat_p1_y ehrnat_p1_z.X : Type a0, a1, a2, b0, b1, b2, c0, c1, c2 : X wry : a0 = b0 wrz : b0 = c0 wly : a1 = b1 wlz : b1 = c1 y : a2 = b2 z : b2 = c2 wryz : a0 = c0 wlyz : a1 = c1 a01 : a0 = a1 a12 : a1 = a2 b01 : b0 = b1 b12 : b1 = b2 c01 : c0 = c1 c12 : c1 = c2 a02 : a0 = a2 c02 : c0 = c2 ehrnat_y : wry @ b01 = a01 @ wly ehrnat_z : wrz @ c01 = b01 @ wlz ulnat_y : wly @ b12 = a12 @ y ulnat_z : wlz @ c12 = b12 @ z wrpp_yz : wry @ wrz = wryz wlpp_yz : wly @ wlz = wlyz H_a02 : a01 @ a12 = a02 H_c02 : c01 @ c12 = c02
((whiskerR wrpp_yz c01)^ @
((ehrnat_y [-] ehrnat_z) @ whiskerL a01 wlpp_yz)
[I] (whiskerR wlpp_yz c12)^ @ (ulnat_y [-] ulnat_z)) @
whiskerR H_a02 (y @ z) =
whiskerL wryz H_c02 @
((whiskerR wrpp_yz c02)^ @
(1 ^ @ ((ehrnat_y [I] ulnat_y) @ whiskerR H_a02 y)
[-] (whiskerL wrz H_c02)^ @
((ehrnat_z [I] ulnat_z) @ 1 )))
destruct H_a02, H_c02.X : Type a0, a1, a2, b0, b1, b2, c0, c1, c2 : X wry : a0 = b0 wrz : b0 = c0 wly : a1 = b1 wlz : b1 = c1 y : a2 = b2 z : b2 = c2 wryz : a0 = c0 wlyz : a1 = c1 a01 : a0 = a1 a12 : a1 = a2 b01 : b0 = b1 b12 : b1 = b2 c01 : c0 = c1 c12 : c1 = c2 ehrnat_y : wry @ b01 = a01 @ wly ehrnat_z : wrz @ c01 = b01 @ wlz ulnat_y : wly @ b12 = a12 @ y ulnat_z : wlz @ c12 = b12 @ z wrpp_yz : wry @ wrz = wryz wlpp_yz : wly @ wlz = wlyz H_a02 : a01 @ a12 = a01 @ a12 H_c02 : c01 @ c12 = c01 @ c12
((whiskerR wrpp_yz c01)^ @
((ehrnat_y [-] ehrnat_z) @ whiskerL a01 wlpp_yz)
[I] (whiskerR wlpp_yz c12)^ @ (ulnat_y [-] ulnat_z)) @
whiskerR 1 (y @ z) =
whiskerL wryz 1 @
((whiskerR wrpp_yz (c01 @ c12))^ @
(1 ^ @ ((ehrnat_y [I] ulnat_y) @ whiskerR 1 y)
[-] (whiskerL wrz 1 )^ @ ((ehrnat_z [I] ulnat_z) @ 1 )))
clear H_a02 H_c02.X : Type a0, a1, a2, b0, b1, b2, c0, c1, c2 : X wry : a0 = b0 wrz : b0 = c0 wly : a1 = b1 wlz : b1 = c1 y : a2 = b2 z : b2 = c2 wryz : a0 = c0 wlyz : a1 = c1 a01 : a0 = a1 a12 : a1 = a2 b01 : b0 = b1 b12 : b1 = b2 c01 : c0 = c1 c12 : c1 = c2 ehrnat_y : wry @ b01 = a01 @ wly ehrnat_z : wrz @ c01 = b01 @ wlz ulnat_y : wly @ b12 = a12 @ y ulnat_z : wlz @ c12 = b12 @ z wrpp_yz : wry @ wrz = wryz wlpp_yz : wly @ wlz = wlyz
((whiskerR wrpp_yz c01)^ @
((ehrnat_y [-] ehrnat_z) @ whiskerL a01 wlpp_yz)
[I] (whiskerR wlpp_yz c12)^ @ (ulnat_y [-] ulnat_z)) @
whiskerR 1 (y @ z) =
whiskerL wryz 1 @
((whiskerR wrpp_yz (c01 @ c12))^ @
(1 ^ @ ((ehrnat_y [I] ulnat_y) @ whiskerR 1 y)
[-] (whiskerL wrz 1 )^ @ ((ehrnat_z [I] ulnat_z) @ 1 )))
destruct wrpp_yz, wlpp_yz.X : Type a0, a1, a2, b0, b1, b2, c0, c1, c2 : X wry : a0 = b0 wrz : b0 = c0 wly : a1 = b1 wlz : b1 = c1 y : a2 = b2 z : b2 = c2 a01 : a0 = a1 a12 : a1 = a2 b01 : b0 = b1 b12 : b1 = b2 c01 : c0 = c1 c12 : c1 = c2 ehrnat_y : wry @ b01 = a01 @ wly ehrnat_z : wrz @ c01 = b01 @ wlz ulnat_y : wly @ b12 = a12 @ y ulnat_z : wlz @ c12 = b12 @ z wrpp_yz : wry @ wrz = wry @ wrz wlpp_yz : wly @ wlz = wly @ wlz
((whiskerR 1 c01)^ @
((ehrnat_y [-] ehrnat_z) @ whiskerL a01 1 )
[I] (whiskerR 1 c12)^ @ (ulnat_y [-] ulnat_z)) @
whiskerR 1 (y @ z) =
whiskerL (wry @ wrz) 1 @
((whiskerR 1 (c01 @ c12))^ @
(1 ^ @ ((ehrnat_y [I] ulnat_y) @ whiskerR 1 y)
[-] (whiskerL wrz 1 )^ @ ((ehrnat_z [I] ulnat_z) @ 1 )))
clear wrpp_yz wlpp_yz.X : Type a0, a1, a2, b0, b1, b2, c0, c1, c2 : X wry : a0 = b0 wrz : b0 = c0 wly : a1 = b1 wlz : b1 = c1 y : a2 = b2 z : b2 = c2 a01 : a0 = a1 a12 : a1 = a2 b01 : b0 = b1 b12 : b1 = b2 c01 : c0 = c1 c12 : c1 = c2 ehrnat_y : wry @ b01 = a01 @ wly ehrnat_z : wrz @ c01 = b01 @ wlz ulnat_y : wly @ b12 = a12 @ y ulnat_z : wlz @ c12 = b12 @ z
((whiskerR 1 c01)^ @
((ehrnat_y [-] ehrnat_z) @ whiskerL a01 1 )
[I] (whiskerR 1 c12)^ @ (ulnat_y [-] ulnat_z)) @
whiskerR 1 (y @ z) =
whiskerL (wry @ wrz) 1 @
((whiskerR 1 (c01 @ c12))^ @
(1 ^ @ ((ehrnat_y [I] ulnat_y) @ whiskerR 1 y)
[-] (whiskerL wrz 1 )^ @ ((ehrnat_z [I] ulnat_z) @ 1 )))
destruct a01, a12, b01, b12, c01, c12.X : Type a0, b0, c0 : X wry : a0 = b0 wrz : b0 = c0 wly : a0 = b0 wlz : b0 = c0 y : a0 = b0 z : b0 = c0 a01, a12 : a0 = a0 b01, b12 : b0 = b0 c01, c12 : c0 = c0 ehrnat_y : wry @ 1 = 1 @ wly ehrnat_z : wrz @ 1 = 1 @ wlz ulnat_y : wly @ 1 = 1 @ y ulnat_z : wlz @ 1 = 1 @ z
((whiskerR 1 1 )^ @
((ehrnat_y [-] ehrnat_z) @ whiskerL 1 1 )
[I] (whiskerR 1 1 )^ @ (ulnat_y [-] ulnat_z)) @
whiskerR 1 (y @ z) =
whiskerL (wry @ wrz) 1 @
((whiskerR 1 (1 @ 1 ))^ @
(1 ^ @ ((ehrnat_y [I] ulnat_y) @ whiskerR 1 y)
[-] (whiskerL wrz 1 )^ @ ((ehrnat_z [I] ulnat_z) @ 1 )))
clear a01 a12 b01 b12 c01 c12.X : Type a0, b0, c0 : X wry : a0 = b0 wrz : b0 = c0 wly : a0 = b0 wlz : b0 = c0 y : a0 = b0 z : b0 = c0 ehrnat_y : wry @ 1 = 1 @ wly ehrnat_z : wrz @ 1 = 1 @ wlz ulnat_y : wly @ 1 = 1 @ y ulnat_z : wlz @ 1 = 1 @ z
((whiskerR 1 1 )^ @
((ehrnat_y [-] ehrnat_z) @ whiskerL 1 1 )
[I] (whiskerR 1 1 )^ @ (ulnat_y [-] ulnat_z)) @
whiskerR 1 (y @ z) =
whiskerL (wry @ wrz) 1 @
((whiskerR 1 (1 @ 1 ))^ @
(1 ^ @ ((ehrnat_y [I] ulnat_y) @ whiskerR 1 y)
[-] (whiskerL wrz 1 )^ @ ((ehrnat_z [I] ulnat_z) @ 1 )))
revert y ulnat_y.X : Type a0, b0, c0 : X wry : a0 = b0 wrz : b0 = c0 wly : a0 = b0 wlz, z : b0 = c0 ehrnat_y : wry @ 1 = 1 @ wly ehrnat_z : wrz @ 1 = 1 @ wlz ulnat_z : wlz @ 1 = 1 @ z
forall (y : a0 = b0) (ulnat_y : wly @ 1 = 1 @ y),
((whiskerR 1 1 )^ @
((ehrnat_y [-] ehrnat_z) @ whiskerL 1 1 )
[I] (whiskerR 1 1 )^ @ (ulnat_y [-] ulnat_z)) @
whiskerR 1 (y @ z) =
whiskerL (wry @ wrz) 1 @
((whiskerR 1 (1 @ 1 ))^ @
(1 ^ @ ((ehrnat_y [I] ulnat_y) @ whiskerR 1 y)
[-] (whiskerL wrz 1 )^ @ ((ehrnat_z [I] ulnat_z) @ 1 )))
snapply equiv_path_ind_rlucancel. X : Type a0, b0, c0 : X wry : a0 = b0 wrz : b0 = c0 wly : a0 = b0 wlz, z : b0 = c0 ehrnat_y : wry @ 1 = 1 @ wly ehrnat_z : wrz @ 1 = 1 @ wlz ulnat_z : wlz @ 1 = 1 @ z
(fun (q : a0 = b0) (s : wly @ 1 = 1 @ q) =>
((whiskerR 1 1 )^ @
((ehrnat_y [-] ehrnat_z) @ whiskerL 1 1 )
[I] (whiskerR 1 1 )^ @ (s [-] ulnat_z)) @
whiskerR 1 (q @ z) =
whiskerL (wry @ wrz) 1 @
((whiskerR 1 (1 @ 1 ))^ @
(1 ^ @ ((ehrnat_y [I] s) @ whiskerR 1 q)
[-] (whiskerL wrz 1 )^ @
((ehrnat_z [I] ulnat_z) @ 1 )))) wly
(rlucancel 1 )
revert z ulnat_z.X : Type a0, b0, c0 : X wry : a0 = b0 wrz : b0 = c0 wly : a0 = b0 wlz : b0 = c0 ehrnat_y : wry @ 1 = 1 @ wly ehrnat_z : wrz @ 1 = 1 @ wlz
forall (z : b0 = c0) (ulnat_z : wlz @ 1 = 1 @ z),
(fun (q : a0 = b0) (s : wly @ 1 = 1 @ q) =>
((whiskerR 1 1 )^ @
((ehrnat_y [-] ehrnat_z) @ whiskerL 1 1 )
[I] (whiskerR 1 1 )^ @ (s [-] ulnat_z)) @
whiskerR 1 (q @ z) =
whiskerL (wry @ wrz) 1 @
((whiskerR 1 (1 @ 1 ))^ @
(1 ^ @ ((ehrnat_y [I] s) @ whiskerR 1 q)
[-] (whiskerL wrz 1 )^ @
((ehrnat_z [I] ulnat_z) @ 1 )))) wly
(rlucancel 1 )
snapply equiv_path_ind_rlucancel. X : Type a0, b0, c0 : X wry : a0 = b0 wrz : b0 = c0 wly : a0 = b0 wlz : b0 = c0 ehrnat_y : wry @ 1 = 1 @ wly ehrnat_z : wrz @ 1 = 1 @ wlz
(fun (q : b0 = c0) (s : wlz @ 1 = 1 @ q) =>
(fun (q0 : a0 = b0) (s0 : wly @ 1 = 1 @ q0) =>
((whiskerR 1 1 )^ @
((ehrnat_y [-] ehrnat_z) @ whiskerL 1 1 )
[I] (whiskerR 1 1 )^ @ (s0 [-] s)) @
whiskerR 1 (q0 @ q) =
whiskerL (wry @ wrz) 1 @
((whiskerR 1 (1 @ 1 ))^ @
(1 ^ @ ((ehrnat_y [I] s0) @ whiskerR 1 q0)
[-] (whiskerL wrz 1 )^ @ ((ehrnat_z [I] s) @ 1 ))))
wly (rlucancel 1 )) wlz (rlucancel 1 )
revert wly ehrnat_y.X : Type a0, b0, c0 : X wry : a0 = b0 wrz, wlz : b0 = c0 ehrnat_z : wrz @ 1 = 1 @ wlz
forall (wly : a0 = b0) (ehrnat_y : wry @ 1 = 1 @ wly),
(fun (q : b0 = c0) (s : wlz @ 1 = 1 @ q) =>
(fun (q0 : a0 = b0) (s0 : wly @ 1 = 1 @ q0) =>
((whiskerR 1 1 )^ @
((ehrnat_y [-] ehrnat_z) @ whiskerL 1 1 )
[I] (whiskerR 1 1 )^ @ (s0 [-] s)) @
whiskerR 1 (q0 @ q) =
whiskerL (wry @ wrz) 1 @
((whiskerR 1 (1 @ 1 ))^ @
(1 ^ @ ((ehrnat_y [I] s0) @ whiskerR 1 q0)
[-] (whiskerL wrz 1 )^ @ ((ehrnat_z [I] s) @ 1 ))))
wly (rlucancel 1 )) wlz (rlucancel 1 )
snapply equiv_path_ind_rlucancel. X : Type a0, b0, c0 : X wry : a0 = b0 wrz, wlz : b0 = c0 ehrnat_z : wrz @ 1 = 1 @ wlz
(fun (q : a0 = b0) (s : wry @ 1 = 1 @ q) =>
(fun (q0 : b0 = c0) (s0 : wlz @ 1 = 1 @ q0) =>
(fun (q1 : a0 = b0) (s1 : q @ 1 = 1 @ q1) =>
((whiskerR 1 1 )^ @
((s [-] ehrnat_z) @ whiskerL 1 1 )
[I] (whiskerR 1 1 )^ @ (s1 [-] s0)) @
whiskerR 1 (q1 @ q0) =
whiskerL (wry @ wrz) 1 @
((whiskerR 1 (1 @ 1 ))^ @
(1 ^ @ ((s [I] s1) @ whiskerR 1 q1)
[-] (whiskerL wrz 1 )^ @ ((ehrnat_z [I] s0) @ 1 ))))
q (rlucancel 1 )) wlz (rlucancel 1 )) wry
(rlucancel 1 )
revert wlz ehrnat_z.X : Type a0, b0, c0 : X wry : a0 = b0 wrz : b0 = c0
forall (wlz : b0 = c0) (ehrnat_z : wrz @ 1 = 1 @ wlz),
(fun (q : a0 = b0) (s : wry @ 1 = 1 @ q) =>
(fun (q0 : b0 = c0) (s0 : wlz @ 1 = 1 @ q0) =>
(fun (q1 : a0 = b0) (s1 : q @ 1 = 1 @ q1) =>
((whiskerR 1 1 )^ @
((s [-] ehrnat_z) @ whiskerL 1 1 )
[I] (whiskerR 1 1 )^ @ (s1 [-] s0)) @
whiskerR 1 (q1 @ q0) =
whiskerL (wry @ wrz) 1 @
((whiskerR 1 (1 @ 1 ))^ @
(1 ^ @ ((s [I] s1) @ whiskerR 1 q1)
[-] (whiskerL wrz 1 )^ @ ((ehrnat_z [I] s0) @ 1 ))))
q (rlucancel 1 )) wlz (rlucancel 1 )) wry
(rlucancel 1 )
snapply equiv_path_ind_rlucancel. X : Type a0, b0, c0 : X wry : a0 = b0 wrz : b0 = c0
(fun (q : b0 = c0) (s : wrz @ 1 = 1 @ q) =>
(fun (q0 : a0 = b0) (s0 : wry @ 1 = 1 @ q0) =>
(fun (q1 : b0 = c0) (s1 : q @ 1 = 1 @ q1) =>
(fun (q2 : a0 = b0) (s2 : q0 @ 1 = 1 @ q2) =>
((whiskerR 1 1 )^ @ ((s0 [-] s) @ whiskerL 1 1 )
[I] (whiskerR 1 1 )^ @ (s2 [-] s1)) @
whiskerR 1 (q2 @ q1) =
whiskerL (wry @ wrz) 1 @
((whiskerR 1 (1 @ 1 ))^ @
(1 ^ @ ((s0 [I] s2) @ whiskerR 1 q2)
[-] (whiskerL wrz 1 )^ @ ((s [I] s1) @ 1 )))) q0
(rlucancel 1 )) q (rlucancel 1 )) wry (rlucancel 1 ))
wrz (rlucancel 1 )
destruct wry, wrz.X : Type a0 : X wry, wrz : a0 = a0
((whiskerR 1 1 )^ @
((rlucancel 1 [-] rlucancel 1 ) @ whiskerL 1 1 )
[I] (whiskerR 1 1 )^ @ (rlucancel 1 [-] rlucancel 1 )) @
whiskerR 1 (1 @ 1 ) =
whiskerL (1 @ 1 ) 1 @
((whiskerR 1 (1 @ 1 ))^ @
(1 ^ @ ((rlucancel 1 [I] rlucancel 1 ) @ whiskerR 1 1 )
[-] (whiskerL 1 1 )^ @
((rlucancel 1 [I] rlucancel 1 ) @ 1 )))
clear wry wrz.X : Type a0 : X
((whiskerR 1 1 )^ @
((rlucancel 1 [-] rlucancel 1 ) @ whiskerL 1 1 )
[I] (whiskerR 1 1 )^ @ (rlucancel 1 [-] rlucancel 1 )) @
whiskerR 1 (1 @ 1 ) =
whiskerL (1 @ 1 ) 1 @
((whiskerR 1 (1 @ 1 ))^ @
(1 ^ @ ((rlucancel 1 [I] rlucancel 1 ) @ whiskerR 1 1 )
[-] (whiskerL 1 1 )^ @
((rlucancel 1 [I] rlucancel 1 ) @ 1 )))
reflexivity .
Defined .
End Ehrnat_p1_pp .
Definition ehrnat_p1_pp {X } {a : X} {u v : idpath a = idpath a} (q : u = 1 ) (r : 1 = v) :
Ehrnat_p1_pp (eh_p1 u) (eh_p1 v) (ehrnat_pp q r 1 ) (ulnat_pp q r) (urnat_pp q r)
(ehrnat_p1 q) (ehrnat_p1 r) =
ehrnat_p1 (q @ r).X : Type a : X u, v : 1 = 1 q : u = 1 r : 1 = v
Ehrnat_p1_pp (eh_p1 u) (eh_p1 v) (ehrnat_pp q r 1 )
(ulnat_pp q r) (urnat_pp q r) (ehrnat_p1 q)
(ehrnat_p1 r) = ehrnat_p1 (q @ r)
Proof .X : Type a : X u, v : 1 = 1 q : u = 1 r : 1 = v
Ehrnat_p1_pp (eh_p1 u) (eh_p1 v) (ehrnat_pp q r 1 )
(ulnat_pp q r) (urnat_pp q r) (ehrnat_p1 q)
(ehrnat_p1 r) = ehrnat_p1 (q @ r)
revert u q.X : Type a : X v : 1 = 1 r : 1 = v
forall (u : 1 = 1 ) (q : u = 1 ),
Ehrnat_p1_pp (eh_p1 u) (eh_p1 v) (ehrnat_pp q r 1 )
(ulnat_pp q r) (urnat_pp q r) (ehrnat_p1 q)
(ehrnat_p1 r) = ehrnat_p1 (q @ r)
snapply (equiv_path_ind (equiv_path_inverse _)). X : Type a : X v : 1 = 1 r : 1 = v
(fun (b : 1 = 1 ) (x : (fun y : 1 = 1 => y = 1 ) b) =>
Ehrnat_p1_pp (eh_p1 b) (eh_p1 v) (ehrnat_pp x r 1 )
(ulnat_pp x r) (urnat_pp x r) (ehrnat_p1 x)
(ehrnat_p1 r) = ehrnat_p1 (x @ r)) 1
(equiv_path_inverse 1 1 1 )
by destruct r.
Defined .
(** Given [wlrnat_V x y] and [wlrnat_V x z], we can explicitly construct [wlrnat_V x (y @ z)]. *)
Section wlrnat_V_p_pp .
Context {X : Type }.
(** 0-paths *)
Context {a0 b0 c0 d0 e0 f0 : X}.
Context {a1 b1 c1 d1 e1 f1 : X}.
(** 1-paths *)
Context {wlx0 : a0 = b0}.
Context {wlx1 : c0 = d0}.
Context {wlx2 : e0 = f0}.
Context {wrx0 : a1 = b1}.
Context {wrx1 : c1 = d1}.
Context {wrx2 : e1 = f1}.
Context {wry0 : b0 = d0}.
Context {wly0 : b1 = d1}.
Context {wry1 : a0 = c0}.
Context {wly1 : a1 = c1}.
Context {wrz0 : d0 = f0}.
Context {wlz0 : d1 = f1}.
Context {wrz1 : c0 = e0}.
Context {wlz1 : c1 = e1}.
Context {a01 : a0 = a1}.
Context {b01 : b0 = b1}.
Context {c01 : c0 = c1}.
Context {d01 : d0 = d1}.
Context {e01 : e0 = e1}.
Context {f01 : f0 = f1}.
Context {wryz0 : b0 = f0}.
Context {wlyz0 : b1 = f1}.
Context {wryz1 : a0 = e0}.
Context {wlyz1 : a1 = e1}.
(** 2-paths *)
Context {ehlnat_x0 : wlx0 @ b01 = a01 @ wrx0}.
Context {ehlnat_x1 : wlx1 @ d01 = c01 @ wrx1}.
Context {ehlnat_x2 : wlx2 @ f01 = e01 @ wrx2}.
Context {ehrnat_y0 : wry0 @ d01 = b01 @ wly0}.
Context {ehrnat_y1 : wry1 @ c01 = a01 @ wly1}.
Context {ehrnat_z0 : wrz0 @ f01 = d01 @ wlz0}.
Context {ehrnat_z1 : wrz1 @ e01 = c01 @ wlz1}.
Context {ehrnat_yz0 : wryz0 @ f01 = b01 @ wlyz0}.
Context {ehrnat_yz1 : wryz1 @ e01 = a01 @ wlyz1}.
Context {wlrnat_x_y : wlx0 @ wry0 = wry1 @ wlx1}.
Context {wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0}.
Context {wlrnat_x_z : wlx1 @ wrz0 = wrz1 @ wlx2}.
Context {wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0}.
Context {wlrnat_x_yz : wlx0 @ wryz0 = wryz1 @ wlx2}.
Context {wlrnat_yz_x : wlyz1 @ wrx2 = wrx0 @ wlyz0}.
Context {wrpp_yz0 : wry0 @ wrz0 = wryz0}.
Context {wlpp_yz0 : wly0 @ wlz0 = wlyz0}.
Context {wrpp_yz1 : wry1 @ wrz1 = wryz1}.
Context {wlpp_yz1 : wly1 @ wlz1 = wlyz1}.
(** 3-paths *)
Hypothesis H_ehrnat_yz0 :
(ehrnat_y0 [-] ehrnat_z0) @ whiskerL _ wlpp_yz0 =
whiskerR wrpp_yz0 _ @ ehrnat_yz0.
Hypothesis H_ehrnat_yz1 :
(ehrnat_y1 [-] ehrnat_z1) @ whiskerL _ wlpp_yz1 =
whiskerR wrpp_yz1 _ @ ehrnat_yz1.
Hypothesis H_wlrnat_x_yz :
(wlrnat_x_y [I] wlrnat_x_z) @ whiskerR wrpp_yz1 _ =
whiskerL _ wrpp_yz0 @ wlrnat_x_yz.
Hypothesis H_wlrnat_yz_x :
(wlrnat_y_x [-] wlrnat_z_x) @ whiskerL _ wlpp_yz0 =
whiskerR wlpp_yz1 _ @ wlrnat_yz_x.
Variable wlrnat_V_x_y :
whiskerR wlrnat_x_y _ @ (ehrnat_y1 [-] ehlnat_x1) =
(ehlnat_x0 [-] ehrnat_y0) @ whiskerL _ wlrnat_y_x^.
Variable wlrnat_V_x_z :
whiskerR wlrnat_x_z _ @ (ehrnat_z1 [-] ehlnat_x2) =
(ehlnat_x1 [-] ehrnat_z0) @ whiskerL _ wlrnat_z_x^.
(** The composite square *)
Definition Wlrnat_V_p_pp :
whiskerR wlrnat_x_yz _ @ (ehrnat_yz1 [-] ehlnat_x2) =
(ehlnat_x0 [-] ehrnat_yz0) @ whiskerL _ wlrnat_yz_x^.X : Type a0, b0, c0, d0, e0, f0, a1, b1, c1, d1, e1, f1 : X wlx0 : a0 = b0 wlx1 : c0 = d0 wlx2 : e0 = f0 wrx0 : a1 = b1 wrx1 : c1 = d1 wrx2 : e1 = f1 wry0 : b0 = d0 wly0 : b1 = d1 wry1 : a0 = c0 wly1 : a1 = c1 wrz0 : d0 = f0 wlz0 : d1 = f1 wrz1 : c0 = e0 wlz1 : c1 = e1 a01 : a0 = a1 b01 : b0 = b1 c01 : c0 = c1 d01 : d0 = d1 e01 : e0 = e1 f01 : f0 = f1 wryz0 : b0 = f0 wlyz0 : b1 = f1 wryz1 : a0 = e0 wlyz1 : a1 = e1 ehlnat_x0 : wlx0 @ b01 = a01 @ wrx0 ehlnat_x1 : wlx1 @ d01 = c01 @ wrx1 ehlnat_x2 : wlx2 @ f01 = e01 @ wrx2 ehrnat_y0 : wry0 @ d01 = b01 @ wly0 ehrnat_y1 : wry1 @ c01 = a01 @ wly1 ehrnat_z0 : wrz0 @ f01 = d01 @ wlz0 ehrnat_z1 : wrz1 @ e01 = c01 @ wlz1 ehrnat_yz0 : wryz0 @ f01 = b01 @ wlyz0 ehrnat_yz1 : wryz1 @ e01 = a01 @ wlyz1 wlrnat_x_y : wlx0 @ wry0 = wry1 @ wlx1 wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_x_z : wlx1 @ wrz0 = wrz1 @ wlx2 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0 wlrnat_x_yz : wlx0 @ wryz0 = wryz1 @ wlx2 wlrnat_yz_x : wlyz1 @ wrx2 = wrx0 @ wlyz0 wrpp_yz0 : wry0 @ wrz0 = wryz0 wlpp_yz0 : wly0 @ wlz0 = wlyz0 wrpp_yz1 : wry1 @ wrz1 = wryz1 wlpp_yz1 : wly1 @ wlz1 = wlyz1 H_ehrnat_yz0 : (ehrnat_y0 [-] ehrnat_z0) @
whiskerL b01 wlpp_yz0 =
whiskerR wrpp_yz0 f01 @ ehrnat_yz0 H_ehrnat_yz1 : (ehrnat_y1 [-] ehrnat_z1) @
whiskerL a01 wlpp_yz1 =
whiskerR wrpp_yz1 e01 @ ehrnat_yz1 H_wlrnat_x_yz : (wlrnat_x_y [I] wlrnat_x_z) @
whiskerR wrpp_yz1 wlx2 =
whiskerL wlx0 wrpp_yz0 @ wlrnat_x_yz H_wlrnat_yz_x : (wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 wlpp_yz0 =
whiskerR wlpp_yz1 wrx2 @ wlrnat_yz_x wlrnat_V_x_y : whiskerR wlrnat_x_y d01 @
(ehrnat_y1 [-] ehlnat_x1) =
(ehlnat_x0 [-] ehrnat_y0) @
whiskerL a01 wlrnat_y_x^ wlrnat_V_x_z : whiskerR wlrnat_x_z f01 @
(ehrnat_z1 [-] ehlnat_x2) =
(ehlnat_x1 [-] ehrnat_z0) @
whiskerL c01 wlrnat_z_x^
whiskerR wlrnat_x_yz f01 @ (ehrnat_yz1 [-] ehlnat_x2) =
(ehlnat_x0 [-] ehrnat_yz0) @ whiskerL a01 wlrnat_yz_x^
Proof .X : Type a0, b0, c0, d0, e0, f0, a1, b1, c1, d1, e1, f1 : X wlx0 : a0 = b0 wlx1 : c0 = d0 wlx2 : e0 = f0 wrx0 : a1 = b1 wrx1 : c1 = d1 wrx2 : e1 = f1 wry0 : b0 = d0 wly0 : b1 = d1 wry1 : a0 = c0 wly1 : a1 = c1 wrz0 : d0 = f0 wlz0 : d1 = f1 wrz1 : c0 = e0 wlz1 : c1 = e1 a01 : a0 = a1 b01 : b0 = b1 c01 : c0 = c1 d01 : d0 = d1 e01 : e0 = e1 f01 : f0 = f1 wryz0 : b0 = f0 wlyz0 : b1 = f1 wryz1 : a0 = e0 wlyz1 : a1 = e1 ehlnat_x0 : wlx0 @ b01 = a01 @ wrx0 ehlnat_x1 : wlx1 @ d01 = c01 @ wrx1 ehlnat_x2 : wlx2 @ f01 = e01 @ wrx2 ehrnat_y0 : wry0 @ d01 = b01 @ wly0 ehrnat_y1 : wry1 @ c01 = a01 @ wly1 ehrnat_z0 : wrz0 @ f01 = d01 @ wlz0 ehrnat_z1 : wrz1 @ e01 = c01 @ wlz1 ehrnat_yz0 : wryz0 @ f01 = b01 @ wlyz0 ehrnat_yz1 : wryz1 @ e01 = a01 @ wlyz1 wlrnat_x_y : wlx0 @ wry0 = wry1 @ wlx1 wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_x_z : wlx1 @ wrz0 = wrz1 @ wlx2 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0 wlrnat_x_yz : wlx0 @ wryz0 = wryz1 @ wlx2 wlrnat_yz_x : wlyz1 @ wrx2 = wrx0 @ wlyz0 wrpp_yz0 : wry0 @ wrz0 = wryz0 wlpp_yz0 : wly0 @ wlz0 = wlyz0 wrpp_yz1 : wry1 @ wrz1 = wryz1 wlpp_yz1 : wly1 @ wlz1 = wlyz1 H_ehrnat_yz0 : (ehrnat_y0 [-] ehrnat_z0) @
whiskerL b01 wlpp_yz0 =
whiskerR wrpp_yz0 f01 @ ehrnat_yz0 H_ehrnat_yz1 : (ehrnat_y1 [-] ehrnat_z1) @
whiskerL a01 wlpp_yz1 =
whiskerR wrpp_yz1 e01 @ ehrnat_yz1 H_wlrnat_x_yz : (wlrnat_x_y [I] wlrnat_x_z) @
whiskerR wrpp_yz1 wlx2 =
whiskerL wlx0 wrpp_yz0 @ wlrnat_x_yz H_wlrnat_yz_x : (wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 wlpp_yz0 =
whiskerR wlpp_yz1 wrx2 @ wlrnat_yz_x wlrnat_V_x_y : whiskerR wlrnat_x_y d01 @
(ehrnat_y1 [-] ehlnat_x1) =
(ehlnat_x0 [-] ehrnat_y0) @
whiskerL a01 wlrnat_y_x^ wlrnat_V_x_z : whiskerR wlrnat_x_z f01 @
(ehrnat_z1 [-] ehlnat_x2) =
(ehlnat_x1 [-] ehrnat_z0) @
whiskerL c01 wlrnat_z_x^
whiskerR wlrnat_x_yz f01 @ (ehrnat_yz1 [-] ehlnat_x2) =
(ehlnat_x0 [-] ehrnat_yz0) @ whiskerL a01 wlrnat_yz_x^
apply moveR_Vp in H_ehrnat_yz0, H_ehrnat_yz1.X : Type a0, b0, c0, d0, e0, f0, a1, b1, c1, d1, e1, f1 : X wlx0 : a0 = b0 wlx1 : c0 = d0 wlx2 : e0 = f0 wrx0 : a1 = b1 wrx1 : c1 = d1 wrx2 : e1 = f1 wry0 : b0 = d0 wly0 : b1 = d1 wry1 : a0 = c0 wly1 : a1 = c1 wrz0 : d0 = f0 wlz0 : d1 = f1 wrz1 : c0 = e0 wlz1 : c1 = e1 a01 : a0 = a1 b01 : b0 = b1 c01 : c0 = c1 d01 : d0 = d1 e01 : e0 = e1 f01 : f0 = f1 wryz0 : b0 = f0 wlyz0 : b1 = f1 wryz1 : a0 = e0 wlyz1 : a1 = e1 ehlnat_x0 : wlx0 @ b01 = a01 @ wrx0 ehlnat_x1 : wlx1 @ d01 = c01 @ wrx1 ehlnat_x2 : wlx2 @ f01 = e01 @ wrx2 ehrnat_y0 : wry0 @ d01 = b01 @ wly0 ehrnat_y1 : wry1 @ c01 = a01 @ wly1 ehrnat_z0 : wrz0 @ f01 = d01 @ wlz0 ehrnat_z1 : wrz1 @ e01 = c01 @ wlz1 ehrnat_yz0 : wryz0 @ f01 = b01 @ wlyz0 ehrnat_yz1 : wryz1 @ e01 = a01 @ wlyz1 wlrnat_x_y : wlx0 @ wry0 = wry1 @ wlx1 wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_x_z : wlx1 @ wrz0 = wrz1 @ wlx2 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0 wlrnat_x_yz : wlx0 @ wryz0 = wryz1 @ wlx2 wlrnat_yz_x : wlyz1 @ wrx2 = wrx0 @ wlyz0 wrpp_yz0 : wry0 @ wrz0 = wryz0 wlpp_yz0 : wly0 @ wlz0 = wlyz0 wrpp_yz1 : wry1 @ wrz1 = wryz1 wlpp_yz1 : wly1 @ wlz1 = wlyz1 H_ehrnat_yz0 : (whiskerR wrpp_yz0 f01)^ @
((ehrnat_y0 [-] ehrnat_z0) @
whiskerL b01 wlpp_yz0) = ehrnat_yz0 H_ehrnat_yz1 : (whiskerR wrpp_yz1 e01)^ @
((ehrnat_y1 [-] ehrnat_z1) @
whiskerL a01 wlpp_yz1) = ehrnat_yz1 H_wlrnat_x_yz : (wlrnat_x_y [I] wlrnat_x_z) @
whiskerR wrpp_yz1 wlx2 =
whiskerL wlx0 wrpp_yz0 @ wlrnat_x_yz H_wlrnat_yz_x : (wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 wlpp_yz0 =
whiskerR wlpp_yz1 wrx2 @ wlrnat_yz_x wlrnat_V_x_y : whiskerR wlrnat_x_y d01 @
(ehrnat_y1 [-] ehlnat_x1) =
(ehlnat_x0 [-] ehrnat_y0) @
whiskerL a01 wlrnat_y_x^ wlrnat_V_x_z : whiskerR wlrnat_x_z f01 @
(ehrnat_z1 [-] ehlnat_x2) =
(ehlnat_x1 [-] ehrnat_z0) @
whiskerL c01 wlrnat_z_x^
whiskerR wlrnat_x_yz f01 @ (ehrnat_yz1 [-] ehlnat_x2) =
(ehlnat_x0 [-] ehrnat_yz0) @ whiskerL a01 wlrnat_yz_x^
destruct H_ehrnat_yz0, H_ehrnat_yz1.X : Type a0, b0, c0, d0, e0, f0, a1, b1, c1, d1, e1, f1 : X wlx0 : a0 = b0 wlx1 : c0 = d0 wlx2 : e0 = f0 wrx0 : a1 = b1 wrx1 : c1 = d1 wrx2 : e1 = f1 wry0 : b0 = d0 wly0 : b1 = d1 wry1 : a0 = c0 wly1 : a1 = c1 wrz0 : d0 = f0 wlz0 : d1 = f1 wrz1 : c0 = e0 wlz1 : c1 = e1 a01 : a0 = a1 b01 : b0 = b1 c01 : c0 = c1 d01 : d0 = d1 e01 : e0 = e1 f01 : f0 = f1 wryz0 : b0 = f0 wlyz0 : b1 = f1 wryz1 : a0 = e0 wlyz1 : a1 = e1 ehlnat_x0 : wlx0 @ b01 = a01 @ wrx0 ehlnat_x1 : wlx1 @ d01 = c01 @ wrx1 ehlnat_x2 : wlx2 @ f01 = e01 @ wrx2 ehrnat_y0 : wry0 @ d01 = b01 @ wly0 ehrnat_y1 : wry1 @ c01 = a01 @ wly1 ehrnat_z0 : wrz0 @ f01 = d01 @ wlz0 ehrnat_z1 : wrz1 @ e01 = c01 @ wlz1 wlrnat_x_y : wlx0 @ wry0 = wry1 @ wlx1 wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_x_z : wlx1 @ wrz0 = wrz1 @ wlx2 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0 wlrnat_x_yz : wlx0 @ wryz0 = wryz1 @ wlx2 wlrnat_yz_x : wlyz1 @ wrx2 = wrx0 @ wlyz0 wrpp_yz0 : wry0 @ wrz0 = wryz0 wlpp_yz0 : wly0 @ wlz0 = wlyz0 wrpp_yz1 : wry1 @ wrz1 = wryz1 wlpp_yz1 : wly1 @ wlz1 = wlyz1 H_ehrnat_yz0 : (whiskerR wrpp_yz0 f01)^ @
((ehrnat_y0 [-] ehrnat_z0) @
whiskerL b01 wlpp_yz0) =
(whiskerR wrpp_yz0 f01)^ @
((ehrnat_y0 [-] ehrnat_z0) @
whiskerL b01 wlpp_yz0) H_ehrnat_yz1 : (whiskerR wrpp_yz1 e01)^ @
((ehrnat_y1 [-] ehrnat_z1) @
whiskerL a01 wlpp_yz1) =
(whiskerR wrpp_yz1 e01)^ @
((ehrnat_y1 [-] ehrnat_z1) @
whiskerL a01 wlpp_yz1) H_wlrnat_x_yz : (wlrnat_x_y [I] wlrnat_x_z) @
whiskerR wrpp_yz1 wlx2 =
whiskerL wlx0 wrpp_yz0 @ wlrnat_x_yz H_wlrnat_yz_x : (wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 wlpp_yz0 =
whiskerR wlpp_yz1 wrx2 @ wlrnat_yz_x wlrnat_V_x_y : whiskerR wlrnat_x_y d01 @
(ehrnat_y1 [-] ehlnat_x1) =
(ehlnat_x0 [-] ehrnat_y0) @
whiskerL a01 wlrnat_y_x^ wlrnat_V_x_z : whiskerR wlrnat_x_z f01 @
(ehrnat_z1 [-] ehlnat_x2) =
(ehlnat_x1 [-] ehrnat_z0) @
whiskerL c01 wlrnat_z_x^
whiskerR wlrnat_x_yz f01 @
((whiskerR wrpp_yz1 e01)^ @
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL a01 wlpp_yz1)
[-] ehlnat_x2) =
(ehlnat_x0
[-] (whiskerR wrpp_yz0 f01)^ @
((ehrnat_y0 [-] ehrnat_z0) @
whiskerL b01 wlpp_yz0)) @
whiskerL a01 wlrnat_yz_x^
clear H_ehrnat_yz0 H_ehrnat_yz1.X : Type a0, b0, c0, d0, e0, f0, a1, b1, c1, d1, e1, f1 : X wlx0 : a0 = b0 wlx1 : c0 = d0 wlx2 : e0 = f0 wrx0 : a1 = b1 wrx1 : c1 = d1 wrx2 : e1 = f1 wry0 : b0 = d0 wly0 : b1 = d1 wry1 : a0 = c0 wly1 : a1 = c1 wrz0 : d0 = f0 wlz0 : d1 = f1 wrz1 : c0 = e0 wlz1 : c1 = e1 a01 : a0 = a1 b01 : b0 = b1 c01 : c0 = c1 d01 : d0 = d1 e01 : e0 = e1 f01 : f0 = f1 wryz0 : b0 = f0 wlyz0 : b1 = f1 wryz1 : a0 = e0 wlyz1 : a1 = e1 ehlnat_x0 : wlx0 @ b01 = a01 @ wrx0 ehlnat_x1 : wlx1 @ d01 = c01 @ wrx1 ehlnat_x2 : wlx2 @ f01 = e01 @ wrx2 ehrnat_y0 : wry0 @ d01 = b01 @ wly0 ehrnat_y1 : wry1 @ c01 = a01 @ wly1 ehrnat_z0 : wrz0 @ f01 = d01 @ wlz0 ehrnat_z1 : wrz1 @ e01 = c01 @ wlz1 wlrnat_x_y : wlx0 @ wry0 = wry1 @ wlx1 wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_x_z : wlx1 @ wrz0 = wrz1 @ wlx2 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0 wlrnat_x_yz : wlx0 @ wryz0 = wryz1 @ wlx2 wlrnat_yz_x : wlyz1 @ wrx2 = wrx0 @ wlyz0 wrpp_yz0 : wry0 @ wrz0 = wryz0 wlpp_yz0 : wly0 @ wlz0 = wlyz0 wrpp_yz1 : wry1 @ wrz1 = wryz1 wlpp_yz1 : wly1 @ wlz1 = wlyz1 H_wlrnat_x_yz : (wlrnat_x_y [I] wlrnat_x_z) @
whiskerR wrpp_yz1 wlx2 =
whiskerL wlx0 wrpp_yz0 @ wlrnat_x_yz H_wlrnat_yz_x : (wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 wlpp_yz0 =
whiskerR wlpp_yz1 wrx2 @ wlrnat_yz_x wlrnat_V_x_y : whiskerR wlrnat_x_y d01 @
(ehrnat_y1 [-] ehlnat_x1) =
(ehlnat_x0 [-] ehrnat_y0) @
whiskerL a01 wlrnat_y_x^ wlrnat_V_x_z : whiskerR wlrnat_x_z f01 @
(ehrnat_z1 [-] ehlnat_x2) =
(ehlnat_x1 [-] ehrnat_z0) @
whiskerL c01 wlrnat_z_x^
whiskerR wlrnat_x_yz f01 @
((whiskerR wrpp_yz1 e01)^ @
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL a01 wlpp_yz1)
[-] ehlnat_x2) =
(ehlnat_x0
[-] (whiskerR wrpp_yz0 f01)^ @
((ehrnat_y0 [-] ehrnat_z0) @
whiskerL b01 wlpp_yz0)) @
whiskerL a01 wlrnat_yz_x^
apply moveR_Vp in H_wlrnat_x_yz, H_wlrnat_yz_x.X : Type a0, b0, c0, d0, e0, f0, a1, b1, c1, d1, e1, f1 : X wlx0 : a0 = b0 wlx1 : c0 = d0 wlx2 : e0 = f0 wrx0 : a1 = b1 wrx1 : c1 = d1 wrx2 : e1 = f1 wry0 : b0 = d0 wly0 : b1 = d1 wry1 : a0 = c0 wly1 : a1 = c1 wrz0 : d0 = f0 wlz0 : d1 = f1 wrz1 : c0 = e0 wlz1 : c1 = e1 a01 : a0 = a1 b01 : b0 = b1 c01 : c0 = c1 d01 : d0 = d1 e01 : e0 = e1 f01 : f0 = f1 wryz0 : b0 = f0 wlyz0 : b1 = f1 wryz1 : a0 = e0 wlyz1 : a1 = e1 ehlnat_x0 : wlx0 @ b01 = a01 @ wrx0 ehlnat_x1 : wlx1 @ d01 = c01 @ wrx1 ehlnat_x2 : wlx2 @ f01 = e01 @ wrx2 ehrnat_y0 : wry0 @ d01 = b01 @ wly0 ehrnat_y1 : wry1 @ c01 = a01 @ wly1 ehrnat_z0 : wrz0 @ f01 = d01 @ wlz0 ehrnat_z1 : wrz1 @ e01 = c01 @ wlz1 wlrnat_x_y : wlx0 @ wry0 = wry1 @ wlx1 wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_x_z : wlx1 @ wrz0 = wrz1 @ wlx2 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0 wlrnat_x_yz : wlx0 @ wryz0 = wryz1 @ wlx2 wlrnat_yz_x : wlyz1 @ wrx2 = wrx0 @ wlyz0 wrpp_yz0 : wry0 @ wrz0 = wryz0 wlpp_yz0 : wly0 @ wlz0 = wlyz0 wrpp_yz1 : wry1 @ wrz1 = wryz1 wlpp_yz1 : wly1 @ wlz1 = wlyz1 H_wlrnat_x_yz : (whiskerL wlx0 wrpp_yz0)^ @
((wlrnat_x_y [I] wlrnat_x_z) @
whiskerR wrpp_yz1 wlx2) = wlrnat_x_yz H_wlrnat_yz_x : (whiskerR wlpp_yz1 wrx2)^ @
((wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 wlpp_yz0) = wlrnat_yz_x wlrnat_V_x_y : whiskerR wlrnat_x_y d01 @
(ehrnat_y1 [-] ehlnat_x1) =
(ehlnat_x0 [-] ehrnat_y0) @
whiskerL a01 wlrnat_y_x^ wlrnat_V_x_z : whiskerR wlrnat_x_z f01 @
(ehrnat_z1 [-] ehlnat_x2) =
(ehlnat_x1 [-] ehrnat_z0) @
whiskerL c01 wlrnat_z_x^
whiskerR wlrnat_x_yz f01 @
((whiskerR wrpp_yz1 e01)^ @
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL a01 wlpp_yz1)
[-] ehlnat_x2) =
(ehlnat_x0
[-] (whiskerR wrpp_yz0 f01)^ @
((ehrnat_y0 [-] ehrnat_z0) @
whiskerL b01 wlpp_yz0)) @
whiskerL a01 wlrnat_yz_x^
destruct H_wlrnat_x_yz, H_wlrnat_yz_x.X : Type a0, b0, c0, d0, e0, f0, a1, b1, c1, d1, e1, f1 : X wlx0 : a0 = b0 wlx1 : c0 = d0 wlx2 : e0 = f0 wrx0 : a1 = b1 wrx1 : c1 = d1 wrx2 : e1 = f1 wry0 : b0 = d0 wly0 : b1 = d1 wry1 : a0 = c0 wly1 : a1 = c1 wrz0 : d0 = f0 wlz0 : d1 = f1 wrz1 : c0 = e0 wlz1 : c1 = e1 a01 : a0 = a1 b01 : b0 = b1 c01 : c0 = c1 d01 : d0 = d1 e01 : e0 = e1 f01 : f0 = f1 wryz0 : b0 = f0 wlyz0 : b1 = f1 wryz1 : a0 = e0 wlyz1 : a1 = e1 ehlnat_x0 : wlx0 @ b01 = a01 @ wrx0 ehlnat_x1 : wlx1 @ d01 = c01 @ wrx1 ehlnat_x2 : wlx2 @ f01 = e01 @ wrx2 ehrnat_y0 : wry0 @ d01 = b01 @ wly0 ehrnat_y1 : wry1 @ c01 = a01 @ wly1 ehrnat_z0 : wrz0 @ f01 = d01 @ wlz0 ehrnat_z1 : wrz1 @ e01 = c01 @ wlz1 wlrnat_x_y : wlx0 @ wry0 = wry1 @ wlx1 wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_x_z : wlx1 @ wrz0 = wrz1 @ wlx2 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0 wrpp_yz0 : wry0 @ wrz0 = wryz0 wlpp_yz0 : wly0 @ wlz0 = wlyz0 wrpp_yz1 : wry1 @ wrz1 = wryz1 wlpp_yz1 : wly1 @ wlz1 = wlyz1 H_wlrnat_x_yz : (whiskerL wlx0 wrpp_yz0)^ @
((wlrnat_x_y [I] wlrnat_x_z) @
whiskerR wrpp_yz1 wlx2) =
(whiskerL wlx0 wrpp_yz0)^ @
((wlrnat_x_y [I] wlrnat_x_z) @
whiskerR wrpp_yz1 wlx2) H_wlrnat_yz_x : (whiskerR wlpp_yz1 wrx2)^ @
((wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 wlpp_yz0) =
(whiskerR wlpp_yz1 wrx2)^ @
((wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 wlpp_yz0) wlrnat_V_x_y : whiskerR wlrnat_x_y d01 @
(ehrnat_y1 [-] ehlnat_x1) =
(ehlnat_x0 [-] ehrnat_y0) @
whiskerL a01 wlrnat_y_x^ wlrnat_V_x_z : whiskerR wlrnat_x_z f01 @
(ehrnat_z1 [-] ehlnat_x2) =
(ehlnat_x1 [-] ehrnat_z0) @
whiskerL c01 wlrnat_z_x^
whiskerR
((whiskerL wlx0 wrpp_yz0)^ @
((wlrnat_x_y [I] wlrnat_x_z) @
whiskerR wrpp_yz1 wlx2)) f01 @
((whiskerR wrpp_yz1 e01)^ @
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL a01 wlpp_yz1)
[-] ehlnat_x2) =
(ehlnat_x0
[-] (whiskerR wrpp_yz0 f01)^ @
((ehrnat_y0 [-] ehrnat_z0) @
whiskerL b01 wlpp_yz0)) @
whiskerL a01
((whiskerR wlpp_yz1 wrx2)^ @
((wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 wlpp_yz0))^
clear H_wlrnat_x_yz H_wlrnat_yz_x.X : Type a0, b0, c0, d0, e0, f0, a1, b1, c1, d1, e1, f1 : X wlx0 : a0 = b0 wlx1 : c0 = d0 wlx2 : e0 = f0 wrx0 : a1 = b1 wrx1 : c1 = d1 wrx2 : e1 = f1 wry0 : b0 = d0 wly0 : b1 = d1 wry1 : a0 = c0 wly1 : a1 = c1 wrz0 : d0 = f0 wlz0 : d1 = f1 wrz1 : c0 = e0 wlz1 : c1 = e1 a01 : a0 = a1 b01 : b0 = b1 c01 : c0 = c1 d01 : d0 = d1 e01 : e0 = e1 f01 : f0 = f1 wryz0 : b0 = f0 wlyz0 : b1 = f1 wryz1 : a0 = e0 wlyz1 : a1 = e1 ehlnat_x0 : wlx0 @ b01 = a01 @ wrx0 ehlnat_x1 : wlx1 @ d01 = c01 @ wrx1 ehlnat_x2 : wlx2 @ f01 = e01 @ wrx2 ehrnat_y0 : wry0 @ d01 = b01 @ wly0 ehrnat_y1 : wry1 @ c01 = a01 @ wly1 ehrnat_z0 : wrz0 @ f01 = d01 @ wlz0 ehrnat_z1 : wrz1 @ e01 = c01 @ wlz1 wlrnat_x_y : wlx0 @ wry0 = wry1 @ wlx1 wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_x_z : wlx1 @ wrz0 = wrz1 @ wlx2 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0 wrpp_yz0 : wry0 @ wrz0 = wryz0 wlpp_yz0 : wly0 @ wlz0 = wlyz0 wrpp_yz1 : wry1 @ wrz1 = wryz1 wlpp_yz1 : wly1 @ wlz1 = wlyz1 wlrnat_V_x_y : whiskerR wlrnat_x_y d01 @
(ehrnat_y1 [-] ehlnat_x1) =
(ehlnat_x0 [-] ehrnat_y0) @
whiskerL a01 wlrnat_y_x^ wlrnat_V_x_z : whiskerR wlrnat_x_z f01 @
(ehrnat_z1 [-] ehlnat_x2) =
(ehlnat_x1 [-] ehrnat_z0) @
whiskerL c01 wlrnat_z_x^
whiskerR
((whiskerL wlx0 wrpp_yz0)^ @
((wlrnat_x_y [I] wlrnat_x_z) @
whiskerR wrpp_yz1 wlx2)) f01 @
((whiskerR wrpp_yz1 e01)^ @
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL a01 wlpp_yz1)
[-] ehlnat_x2) =
(ehlnat_x0
[-] (whiskerR wrpp_yz0 f01)^ @
((ehrnat_y0 [-] ehrnat_z0) @
whiskerL b01 wlpp_yz0)) @
whiskerL a01
((whiskerR wlpp_yz1 wrx2)^ @
((wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 wlpp_yz0))^
destruct a01, b01, c01, d01, e01, f01.X : Type a0, b0, c0, d0, e0, f0 : X wlx0 : a0 = b0 wlx1 : c0 = d0 wlx2 : e0 = f0 wrx0 : a0 = b0 wrx1 : c0 = d0 wrx2 : e0 = f0 wry0, wly0 : b0 = d0 wry1, wly1 : a0 = c0 wrz0, wlz0 : d0 = f0 wrz1, wlz1 : c0 = e0 a01 : a0 = a0 b01 : b0 = b0 c01 : c0 = c0 d01 : d0 = d0 e01 : e0 = e0 f01 : f0 = f0 wryz0, wlyz0 : b0 = f0 wryz1, wlyz1 : a0 = e0 ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : wry0 @ 1 = 1 @ wly0 ehrnat_y1 : wry1 @ 1 = 1 @ wly1 ehrnat_z0 : wrz0 @ 1 = 1 @ wlz0 ehrnat_z1 : wrz1 @ 1 = 1 @ wlz1 wlrnat_x_y : wlx0 @ wry0 = wry1 @ wlx1 wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_x_z : wlx1 @ wrz0 = wrz1 @ wlx2 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0 wrpp_yz0 : wry0 @ wrz0 = wryz0 wlpp_yz0 : wly0 @ wlz0 = wlyz0 wrpp_yz1 : wry1 @ wrz1 = wryz1 wlpp_yz1 : wly1 @ wlz1 = wlyz1 wlrnat_V_x_y : whiskerR wlrnat_x_y 1 @
(ehrnat_y1 [-] ehlnat_x1) =
(ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^ wlrnat_V_x_z : whiskerR wlrnat_x_z 1 @
(ehrnat_z1 [-] ehlnat_x2) =
(ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^
whiskerR
((whiskerL wlx0 wrpp_yz0)^ @
((wlrnat_x_y [I] wlrnat_x_z) @
whiskerR wrpp_yz1 wlx2)) 1 @
((whiskerR wrpp_yz1 1 )^ @
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 wlpp_yz1)
[-] ehlnat_x2) =
(ehlnat_x0
[-] (whiskerR wrpp_yz0 1 )^ @
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 wlpp_yz0)) @
whiskerL 1
((whiskerR wlpp_yz1 wrx2)^ @
((wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 wlpp_yz0))^
clear a01 b01 c01 d01 e01 f01.X : Type a0, b0, c0, d0, e0, f0 : X wlx0 : a0 = b0 wlx1 : c0 = d0 wlx2 : e0 = f0 wrx0 : a0 = b0 wrx1 : c0 = d0 wrx2 : e0 = f0 wry0, wly0 : b0 = d0 wry1, wly1 : a0 = c0 wrz0, wlz0 : d0 = f0 wrz1, wlz1 : c0 = e0 wryz0, wlyz0 : b0 = f0 wryz1, wlyz1 : a0 = e0 ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : wry0 @ 1 = 1 @ wly0 ehrnat_y1 : wry1 @ 1 = 1 @ wly1 ehrnat_z0 : wrz0 @ 1 = 1 @ wlz0 ehrnat_z1 : wrz1 @ 1 = 1 @ wlz1 wlrnat_x_y : wlx0 @ wry0 = wry1 @ wlx1 wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_x_z : wlx1 @ wrz0 = wrz1 @ wlx2 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0 wrpp_yz0 : wry0 @ wrz0 = wryz0 wlpp_yz0 : wly0 @ wlz0 = wlyz0 wrpp_yz1 : wry1 @ wrz1 = wryz1 wlpp_yz1 : wly1 @ wlz1 = wlyz1 wlrnat_V_x_y : whiskerR wlrnat_x_y 1 @
(ehrnat_y1 [-] ehlnat_x1) =
(ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^ wlrnat_V_x_z : whiskerR wlrnat_x_z 1 @
(ehrnat_z1 [-] ehlnat_x2) =
(ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^
whiskerR
((whiskerL wlx0 wrpp_yz0)^ @
((wlrnat_x_y [I] wlrnat_x_z) @
whiskerR wrpp_yz1 wlx2)) 1 @
((whiskerR wrpp_yz1 1 )^ @
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 wlpp_yz1)
[-] ehlnat_x2) =
(ehlnat_x0
[-] (whiskerR wrpp_yz0 1 )^ @
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 wlpp_yz0)) @
whiskerL 1
((whiskerR wlpp_yz1 wrx2)^ @
((wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 wlpp_yz0))^
pose (H_whiskerR_wlrnat_x_y := moveL_Mp _ _ _ (moveL_pV _ _ _ (whiskerR_p1 wlrnat_x_y))).X : Type a0, b0, c0, d0, e0, f0 : X wlx0 : a0 = b0 wlx1 : c0 = d0 wlx2 : e0 = f0 wrx0 : a0 = b0 wrx1 : c0 = d0 wrx2 : e0 = f0 wry0, wly0 : b0 = d0 wry1, wly1 : a0 = c0 wrz0, wlz0 : d0 = f0 wrz1, wlz1 : c0 = e0 wryz0, wlyz0 : b0 = f0 wryz1, wlyz1 : a0 = e0 ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : wry0 @ 1 = 1 @ wly0 ehrnat_y1 : wry1 @ 1 = 1 @ wly1 ehrnat_z0 : wrz0 @ 1 = 1 @ wlz0 ehrnat_z1 : wrz1 @ 1 = 1 @ wlz1 wlrnat_x_y : wlx0 @ wry0 = wry1 @ wlx1 wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_x_z : wlx1 @ wrz0 = wrz1 @ wlx2 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0 wrpp_yz0 : wry0 @ wrz0 = wryz0 wlpp_yz0 : wly0 @ wlz0 = wlyz0 wrpp_yz1 : wry1 @ wrz1 = wryz1 wlpp_yz1 : wly1 @ wlz1 = wlyz1 wlrnat_V_x_y : whiskerR wlrnat_x_y 1 @
(ehrnat_y1 [-] ehlnat_x1) =
(ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^ wlrnat_V_x_z : whiskerR wlrnat_x_z 1 @
(ehrnat_z1 [-] ehlnat_x2) =
(ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^ H_whiskerR_wlrnat_x_y := moveL_Mp
(wlrnat_x_y @
(concat_p1 (wry1 @ wlx1))^)
(whiskerR wlrnat_x_y 1 )
(concat_p1 (wlx0 @ wry0))
(moveL_pV
(concat_p1 (wry1 @ wlx1))
((concat_p1 (wlx0 @ wry0))^ @
whiskerR wlrnat_x_y 1 )
wlrnat_x_y
(whiskerR_p1 wlrnat_x_y)) : whiskerR wlrnat_x_y 1 =
concat_p1 (wlx0 @ wry0) @
(wlrnat_x_y @
(concat_p1 (wry1 @ wlx1))^)
whiskerR
((whiskerL wlx0 wrpp_yz0)^ @
((wlrnat_x_y [I] wlrnat_x_z) @
whiskerR wrpp_yz1 wlx2)) 1 @
((whiskerR wrpp_yz1 1 )^ @
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 wlpp_yz1)
[-] ehlnat_x2) =
(ehlnat_x0
[-] (whiskerR wrpp_yz0 1 )^ @
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 wlpp_yz0)) @
whiskerL 1
((whiskerR wlpp_yz1 wrx2)^ @
((wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 wlpp_yz0))^
pose (H_whiskerR_wlrnat_x_z := moveL_Mp _ _ _ (moveL_pV _ _ _ (whiskerR_p1 wlrnat_x_z))).X : Type a0, b0, c0, d0, e0, f0 : X wlx0 : a0 = b0 wlx1 : c0 = d0 wlx2 : e0 = f0 wrx0 : a0 = b0 wrx1 : c0 = d0 wrx2 : e0 = f0 wry0, wly0 : b0 = d0 wry1, wly1 : a0 = c0 wrz0, wlz0 : d0 = f0 wrz1, wlz1 : c0 = e0 wryz0, wlyz0 : b0 = f0 wryz1, wlyz1 : a0 = e0 ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : wry0 @ 1 = 1 @ wly0 ehrnat_y1 : wry1 @ 1 = 1 @ wly1 ehrnat_z0 : wrz0 @ 1 = 1 @ wlz0 ehrnat_z1 : wrz1 @ 1 = 1 @ wlz1 wlrnat_x_y : wlx0 @ wry0 = wry1 @ wlx1 wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_x_z : wlx1 @ wrz0 = wrz1 @ wlx2 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0 wrpp_yz0 : wry0 @ wrz0 = wryz0 wlpp_yz0 : wly0 @ wlz0 = wlyz0 wrpp_yz1 : wry1 @ wrz1 = wryz1 wlpp_yz1 : wly1 @ wlz1 = wlyz1 wlrnat_V_x_y : whiskerR wlrnat_x_y 1 @
(ehrnat_y1 [-] ehlnat_x1) =
(ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^ wlrnat_V_x_z : whiskerR wlrnat_x_z 1 @
(ehrnat_z1 [-] ehlnat_x2) =
(ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^ H_whiskerR_wlrnat_x_y := moveL_Mp
(wlrnat_x_y @
(concat_p1 (wry1 @ wlx1))^)
(whiskerR wlrnat_x_y 1 )
(concat_p1 (wlx0 @ wry0))
(moveL_pV
(concat_p1 (wry1 @ wlx1))
((concat_p1 (wlx0 @ wry0))^ @
whiskerR wlrnat_x_y 1 )
wlrnat_x_y
(whiskerR_p1 wlrnat_x_y)) : whiskerR wlrnat_x_y 1 =
concat_p1 (wlx0 @ wry0) @
(wlrnat_x_y @
(concat_p1 (wry1 @ wlx1))^) H_whiskerR_wlrnat_x_z := moveL_Mp
(wlrnat_x_z @
(concat_p1 (wrz1 @ wlx2))^)
(whiskerR wlrnat_x_z 1 )
(concat_p1 (wlx1 @ wrz0))
(moveL_pV
(concat_p1 (wrz1 @ wlx2))
((concat_p1 (wlx1 @ wrz0))^ @
whiskerR wlrnat_x_z 1 )
wlrnat_x_z
(whiskerR_p1 wlrnat_x_z)) : whiskerR wlrnat_x_z 1 =
concat_p1 (wlx1 @ wrz0) @
(wlrnat_x_z @
(concat_p1 (wrz1 @ wlx2))^)
whiskerR
((whiskerL wlx0 wrpp_yz0)^ @
((wlrnat_x_y [I] wlrnat_x_z) @
whiskerR wrpp_yz1 wlx2)) 1 @
((whiskerR wrpp_yz1 1 )^ @
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 wlpp_yz1)
[-] ehlnat_x2) =
(ehlnat_x0
[-] (whiskerR wrpp_yz0 1 )^ @
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 wlpp_yz0)) @
whiskerL 1
((whiskerR wlpp_yz1 wrx2)^ @
((wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 wlpp_yz0))^
apply moveL_pV in wlrnat_V_x_y.X : Type a0, b0, c0, d0, e0, f0 : X wlx0 : a0 = b0 wlx1 : c0 = d0 wlx2 : e0 = f0 wrx0 : a0 = b0 wrx1 : c0 = d0 wrx2 : e0 = f0 wry0, wly0 : b0 = d0 wry1, wly1 : a0 = c0 wrz0, wlz0 : d0 = f0 wrz1, wlz1 : c0 = e0 wryz0, wlyz0 : b0 = f0 wryz1, wlyz1 : a0 = e0 ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : wry0 @ 1 = 1 @ wly0 ehrnat_y1 : wry1 @ 1 = 1 @ wly1 ehrnat_z0 : wrz0 @ 1 = 1 @ wlz0 ehrnat_z1 : wrz1 @ 1 = 1 @ wlz1 wlrnat_x_y : wlx0 @ wry0 = wry1 @ wlx1 wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_x_z : wlx1 @ wrz0 = wrz1 @ wlx2 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0 wrpp_yz0 : wry0 @ wrz0 = wryz0 wlpp_yz0 : wly0 @ wlz0 = wlyz0 wrpp_yz1 : wry1 @ wrz1 = wryz1 wlpp_yz1 : wly1 @ wlz1 = wlyz1 wlrnat_V_x_y : whiskerR wlrnat_x_y 1 =
((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^ wlrnat_V_x_z : whiskerR wlrnat_x_z 1 @
(ehrnat_z1 [-] ehlnat_x2) =
(ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^ H_whiskerR_wlrnat_x_y := moveL_Mp
(wlrnat_x_y @
(concat_p1 (wry1 @ wlx1))^)
(whiskerR wlrnat_x_y 1 )
(concat_p1 (wlx0 @ wry0))
(moveL_pV
(concat_p1 (wry1 @ wlx1))
((concat_p1 (wlx0 @ wry0))^ @
whiskerR wlrnat_x_y 1 )
wlrnat_x_y
(whiskerR_p1 wlrnat_x_y)) : whiskerR wlrnat_x_y 1 =
concat_p1 (wlx0 @ wry0) @
(wlrnat_x_y @
(concat_p1 (wry1 @ wlx1))^) H_whiskerR_wlrnat_x_z := moveL_Mp
(wlrnat_x_z @
(concat_p1 (wrz1 @ wlx2))^)
(whiskerR wlrnat_x_z 1 )
(concat_p1 (wlx1 @ wrz0))
(moveL_pV
(concat_p1 (wrz1 @ wlx2))
((concat_p1 (wlx1 @ wrz0))^ @
whiskerR wlrnat_x_z 1 )
wlrnat_x_z
(whiskerR_p1 wlrnat_x_z)) : whiskerR wlrnat_x_z 1 =
concat_p1 (wlx1 @ wrz0) @
(wlrnat_x_z @
(concat_p1 (wrz1 @ wlx2))^)
whiskerR
((whiskerL wlx0 wrpp_yz0)^ @
((wlrnat_x_y [I] wlrnat_x_z) @
whiskerR wrpp_yz1 wlx2)) 1 @
((whiskerR wrpp_yz1 1 )^ @
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 wlpp_yz1)
[-] ehlnat_x2) =
(ehlnat_x0
[-] (whiskerR wrpp_yz0 1 )^ @
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 wlpp_yz0)) @
whiskerL 1
((whiskerR wlpp_yz1 wrx2)^ @
((wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 wlpp_yz0))^
apply (concat H_whiskerR_wlrnat_x_y^) in wlrnat_V_x_y.X : Type a0, b0, c0, d0, e0, f0 : X wlx0 : a0 = b0 wlx1 : c0 = d0 wlx2 : e0 = f0 wrx0 : a0 = b0 wrx1 : c0 = d0 wrx2 : e0 = f0 wry0, wly0 : b0 = d0 wry1, wly1 : a0 = c0 wrz0, wlz0 : d0 = f0 wrz1, wlz1 : c0 = e0 wryz0, wlyz0 : b0 = f0 wryz1, wlyz1 : a0 = e0 ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : wry0 @ 1 = 1 @ wly0 ehrnat_y1 : wry1 @ 1 = 1 @ wly1 ehrnat_z0 : wrz0 @ 1 = 1 @ wlz0 ehrnat_z1 : wrz1 @ 1 = 1 @ wlz1 wlrnat_x_y : wlx0 @ wry0 = wry1 @ wlx1 wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_x_z : wlx1 @ wrz0 = wrz1 @ wlx2 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0 wrpp_yz0 : wry0 @ wrz0 = wryz0 wlpp_yz0 : wly0 @ wlz0 = wlyz0 wrpp_yz1 : wry1 @ wrz1 = wryz1 wlpp_yz1 : wly1 @ wlz1 = wlyz1 wlrnat_V_x_y : concat_p1 (wlx0 @ wry0) @
(wlrnat_x_y @
(concat_p1 (wry1 @ wlx1))^) =
((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^ wlrnat_V_x_z : whiskerR wlrnat_x_z 1 @
(ehrnat_z1 [-] ehlnat_x2) =
(ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^ H_whiskerR_wlrnat_x_y := moveL_Mp
(wlrnat_x_y @
(concat_p1 (wry1 @ wlx1))^)
(whiskerR wlrnat_x_y 1 )
(concat_p1 (wlx0 @ wry0))
(moveL_pV
(concat_p1 (wry1 @ wlx1))
((concat_p1 (wlx0 @ wry0))^ @
whiskerR wlrnat_x_y 1 )
wlrnat_x_y
(whiskerR_p1 wlrnat_x_y)) : whiskerR wlrnat_x_y 1 =
concat_p1 (wlx0 @ wry0) @
(wlrnat_x_y @
(concat_p1 (wry1 @ wlx1))^) H_whiskerR_wlrnat_x_z := moveL_Mp
(wlrnat_x_z @
(concat_p1 (wrz1 @ wlx2))^)
(whiskerR wlrnat_x_z 1 )
(concat_p1 (wlx1 @ wrz0))
(moveL_pV
(concat_p1 (wrz1 @ wlx2))
((concat_p1 (wlx1 @ wrz0))^ @
whiskerR wlrnat_x_z 1 )
wlrnat_x_z
(whiskerR_p1 wlrnat_x_z)) : whiskerR wlrnat_x_z 1 =
concat_p1 (wlx1 @ wrz0) @
(wlrnat_x_z @
(concat_p1 (wrz1 @ wlx2))^)
whiskerR
((whiskerL wlx0 wrpp_yz0)^ @
((wlrnat_x_y [I] wlrnat_x_z) @
whiskerR wrpp_yz1 wlx2)) 1 @
((whiskerR wrpp_yz1 1 )^ @
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 wlpp_yz1)
[-] ehlnat_x2) =
(ehlnat_x0
[-] (whiskerR wrpp_yz0 1 )^ @
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 wlpp_yz0)) @
whiskerL 1
((whiskerR wlpp_yz1 wrx2)^ @
((wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 wlpp_yz0))^
apply moveL_Vp, moveL_pV in wlrnat_V_x_y.X : Type a0, b0, c0, d0, e0, f0 : X wlx0 : a0 = b0 wlx1 : c0 = d0 wlx2 : e0 = f0 wrx0 : a0 = b0 wrx1 : c0 = d0 wrx2 : e0 = f0 wry0, wly0 : b0 = d0 wry1, wly1 : a0 = c0 wrz0, wlz0 : d0 = f0 wrz1, wlz1 : c0 = e0 wryz0, wlyz0 : b0 = f0 wryz1, wlyz1 : a0 = e0 ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : wry0 @ 1 = 1 @ wly0 ehrnat_y1 : wry1 @ 1 = 1 @ wly1 ehrnat_z0 : wrz0 @ 1 = 1 @ wlz0 ehrnat_z1 : wrz1 @ 1 = 1 @ wlz1 wlrnat_x_y : wlx0 @ wry0 = wry1 @ wlx1 wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_x_z : wlx1 @ wrz0 = wrz1 @ wlx2 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0 wrpp_yz0 : wry0 @ wrz0 = wryz0 wlpp_yz0 : wly0 @ wlz0 = wlyz0 wrpp_yz1 : wry1 @ wrz1 = wryz1 wlpp_yz1 : wly1 @ wlz1 = wlyz1 wlrnat_V_x_y : wlrnat_x_y =
((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^ wlrnat_V_x_z : whiskerR wlrnat_x_z 1 @
(ehrnat_z1 [-] ehlnat_x2) =
(ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^ H_whiskerR_wlrnat_x_y := moveL_Mp
(wlrnat_x_y @
(concat_p1 (wry1 @ wlx1))^)
(whiskerR wlrnat_x_y 1 )
(concat_p1 (wlx0 @ wry0))
(moveL_pV
(concat_p1 (wry1 @ wlx1))
((concat_p1 (wlx0 @ wry0))^ @
whiskerR wlrnat_x_y 1 )
wlrnat_x_y
(whiskerR_p1 wlrnat_x_y)) : whiskerR wlrnat_x_y 1 =
concat_p1 (wlx0 @ wry0) @
(wlrnat_x_y @
(concat_p1 (wry1 @ wlx1))^) H_whiskerR_wlrnat_x_z := moveL_Mp
(wlrnat_x_z @
(concat_p1 (wrz1 @ wlx2))^)
(whiskerR wlrnat_x_z 1 )
(concat_p1 (wlx1 @ wrz0))
(moveL_pV
(concat_p1 (wrz1 @ wlx2))
((concat_p1 (wlx1 @ wrz0))^ @
whiskerR wlrnat_x_z 1 )
wlrnat_x_z
(whiskerR_p1 wlrnat_x_z)) : whiskerR wlrnat_x_z 1 =
concat_p1 (wlx1 @ wrz0) @
(wlrnat_x_z @
(concat_p1 (wrz1 @ wlx2))^)
whiskerR
((whiskerL wlx0 wrpp_yz0)^ @
((wlrnat_x_y [I] wlrnat_x_z) @
whiskerR wrpp_yz1 wlx2)) 1 @
((whiskerR wrpp_yz1 1 )^ @
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 wlpp_yz1)
[-] ehlnat_x2) =
(ehlnat_x0
[-] (whiskerR wrpp_yz0 1 )^ @
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 wlpp_yz0)) @
whiskerL 1
((whiskerR wlpp_yz1 wrx2)^ @
((wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 wlpp_yz0))^
apply symmetry in wlrnat_V_x_y.X : Type a0, b0, c0, d0, e0, f0 : X wlx0 : a0 = b0 wlx1 : c0 = d0 wlx2 : e0 = f0 wrx0 : a0 = b0 wrx1 : c0 = d0 wrx2 : e0 = f0 wry0, wly0 : b0 = d0 wry1, wly1 : a0 = c0 wrz0, wlz0 : d0 = f0 wrz1, wlz1 : c0 = e0 wryz0, wlyz0 : b0 = f0 wryz1, wlyz1 : a0 = e0 ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : wry0 @ 1 = 1 @ wly0 ehrnat_y1 : wry1 @ 1 = 1 @ wly1 ehrnat_z0 : wrz0 @ 1 = 1 @ wlz0 ehrnat_z1 : wrz1 @ 1 = 1 @ wlz1 wlrnat_x_y : wlx0 @ wry0 = wry1 @ wlx1 wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_x_z : wlx1 @ wrz0 = wrz1 @ wlx2 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0 wrpp_yz0 : wry0 @ wrz0 = wryz0 wlpp_yz0 : wly0 @ wlz0 = wlyz0 wrpp_yz1 : wry1 @ wrz1 = wryz1 wlpp_yz1 : wly1 @ wlz1 = wlyz1 wlrnat_V_x_y : ((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^ =
wlrnat_x_y wlrnat_V_x_z : whiskerR wlrnat_x_z 1 @
(ehrnat_z1 [-] ehlnat_x2) =
(ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^ H_whiskerR_wlrnat_x_y := moveL_Mp
(wlrnat_x_y @
(concat_p1 (wry1 @ wlx1))^)
(whiskerR wlrnat_x_y 1 )
(concat_p1 (wlx0 @ wry0))
(moveL_pV
(concat_p1 (wry1 @ wlx1))
((concat_p1 (wlx0 @ wry0))^ @
whiskerR wlrnat_x_y 1 )
wlrnat_x_y
(whiskerR_p1 wlrnat_x_y)) : whiskerR wlrnat_x_y 1 =
concat_p1 (wlx0 @ wry0) @
(wlrnat_x_y @
(concat_p1 (wry1 @ wlx1))^) H_whiskerR_wlrnat_x_z := moveL_Mp
(wlrnat_x_z @
(concat_p1 (wrz1 @ wlx2))^)
(whiskerR wlrnat_x_z 1 )
(concat_p1 (wlx1 @ wrz0))
(moveL_pV
(concat_p1 (wrz1 @ wlx2))
((concat_p1 (wlx1 @ wrz0))^ @
whiskerR wlrnat_x_z 1 )
wlrnat_x_z
(whiskerR_p1 wlrnat_x_z)) : whiskerR wlrnat_x_z 1 =
concat_p1 (wlx1 @ wrz0) @
(wlrnat_x_z @
(concat_p1 (wrz1 @ wlx2))^)
whiskerR
((whiskerL wlx0 wrpp_yz0)^ @
((wlrnat_x_y [I] wlrnat_x_z) @
whiskerR wrpp_yz1 wlx2)) 1 @
((whiskerR wrpp_yz1 1 )^ @
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 wlpp_yz1)
[-] ehlnat_x2) =
(ehlnat_x0
[-] (whiskerR wrpp_yz0 1 )^ @
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 wlpp_yz0)) @
whiskerL 1
((whiskerR wlpp_yz1 wrx2)^ @
((wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 wlpp_yz0))^
destruct wlrnat_V_x_y.X : Type a0, b0, c0, d0, e0, f0 : X wlx0 : a0 = b0 wlx1 : c0 = d0 wlx2 : e0 = f0 wrx0 : a0 = b0 wrx1 : c0 = d0 wrx2 : e0 = f0 wry0, wly0 : b0 = d0 wry1, wly1 : a0 = c0 wrz0, wlz0 : d0 = f0 wrz1, wlz1 : c0 = e0 wryz0, wlyz0 : b0 = f0 wryz1, wlyz1 : a0 = e0 ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : wry0 @ 1 = 1 @ wly0 ehrnat_y1 : wry1 @ 1 = 1 @ wly1 ehrnat_z0 : wrz0 @ 1 = 1 @ wlz0 ehrnat_z1 : wrz1 @ 1 = 1 @ wlz1 wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_x_z : wlx1 @ wrz0 = wrz1 @ wlx2 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0 wrpp_yz0 : wry0 @ wrz0 = wryz0 wlpp_yz0 : wly0 @ wlz0 = wlyz0 wrpp_yz1 : wry1 @ wrz1 = wryz1 wlpp_yz1 : wly1 @ wlz1 = wlyz1 wlrnat_V_x_y : ((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^ =
((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^ wlrnat_V_x_z : whiskerR wlrnat_x_z 1 @
(ehrnat_z1 [-] ehlnat_x2) =
(ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^ H_whiskerR_wlrnat_x_y := moveL_Mp
((((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^) @
(concat_p1 (wry1 @ wlx1))^)
(whiskerR
(((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0
[-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1
[-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^)
1 )
(concat_p1 (wlx0 @ wry0))
(moveL_pV
(concat_p1 (wry1 @ wlx1))
((concat_p1 (wlx0 @ wry0))^ @
whiskerR
(((concat_p1
(wlx0 @ wry0))^ @
(((ehlnat_x0
[-] ehrnat_y0) @
whiskerL 1
wlrnat_y_x^) @
(ehrnat_y1
[-] ehlnat_x1)^)) @
((concat_p1
(wry1 @ wlx1))^)^)
1 )
(((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0
[-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1
[-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^)
(whiskerR_p1
(((concat_p1
(wlx0 @ wry0))^ @
(((ehlnat_x0
[-] ehrnat_y0) @
whiskerL 1
wlrnat_y_x^) @
(ehrnat_y1
[-] ehlnat_x1)^)) @
((concat_p1
(wry1 @ wlx1))^)^))) : whiskerR
(((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^)
1 =
concat_p1 (wlx0 @ wry0) @
((((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^) @
(concat_p1 (wry1 @ wlx1))^) H_whiskerR_wlrnat_x_z := moveL_Mp
(wlrnat_x_z @
(concat_p1 (wrz1 @ wlx2))^)
(whiskerR wlrnat_x_z 1 )
(concat_p1 (wlx1 @ wrz0))
(moveL_pV
(concat_p1 (wrz1 @ wlx2))
((concat_p1 (wlx1 @ wrz0))^ @
whiskerR wlrnat_x_z 1 )
wlrnat_x_z
(whiskerR_p1 wlrnat_x_z)) : whiskerR wlrnat_x_z 1 =
concat_p1 (wlx1 @ wrz0) @
(wlrnat_x_z @
(concat_p1 (wrz1 @ wlx2))^)
whiskerR
((whiskerL wlx0 wrpp_yz0)^ @
((((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^ [I] wlrnat_x_z) @
whiskerR wrpp_yz1 wlx2)) 1 @
((whiskerR wrpp_yz1 1 )^ @
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 wlpp_yz1)
[-] ehlnat_x2) =
(ehlnat_x0
[-] (whiskerR wrpp_yz0 1 )^ @
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 wlpp_yz0)) @
whiskerL 1
((whiskerR wlpp_yz1 wrx2)^ @
((wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 wlpp_yz0))^
clear wlrnat_V_x_y.X : Type a0, b0, c0, d0, e0, f0 : X wlx0 : a0 = b0 wlx1 : c0 = d0 wlx2 : e0 = f0 wrx0 : a0 = b0 wrx1 : c0 = d0 wrx2 : e0 = f0 wry0, wly0 : b0 = d0 wry1, wly1 : a0 = c0 wrz0, wlz0 : d0 = f0 wrz1, wlz1 : c0 = e0 wryz0, wlyz0 : b0 = f0 wryz1, wlyz1 : a0 = e0 ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : wry0 @ 1 = 1 @ wly0 ehrnat_y1 : wry1 @ 1 = 1 @ wly1 ehrnat_z0 : wrz0 @ 1 = 1 @ wlz0 ehrnat_z1 : wrz1 @ 1 = 1 @ wlz1 wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_x_z : wlx1 @ wrz0 = wrz1 @ wlx2 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0 wrpp_yz0 : wry0 @ wrz0 = wryz0 wlpp_yz0 : wly0 @ wlz0 = wlyz0 wrpp_yz1 : wry1 @ wrz1 = wryz1 wlpp_yz1 : wly1 @ wlz1 = wlyz1 wlrnat_V_x_z : whiskerR wlrnat_x_z 1 @
(ehrnat_z1 [-] ehlnat_x2) =
(ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^ H_whiskerR_wlrnat_x_y := moveL_Mp
((((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^) @
(concat_p1 (wry1 @ wlx1))^)
(whiskerR
(((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0
[-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1
[-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^)
1 )
(concat_p1 (wlx0 @ wry0))
(moveL_pV
(concat_p1 (wry1 @ wlx1))
((concat_p1 (wlx0 @ wry0))^ @
whiskerR
(((concat_p1
(wlx0 @ wry0))^ @
(((ehlnat_x0
[-] ehrnat_y0) @
whiskerL 1
wlrnat_y_x^) @
(ehrnat_y1
[-] ehlnat_x1)^)) @
((concat_p1
(wry1 @ wlx1))^)^)
1 )
(((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0
[-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1
[-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^)
(whiskerR_p1
(((concat_p1
(wlx0 @ wry0))^ @
(((ehlnat_x0
[-] ehrnat_y0) @
whiskerL 1
wlrnat_y_x^) @
(ehrnat_y1
[-] ehlnat_x1)^)) @
((concat_p1
(wry1 @ wlx1))^)^))) : whiskerR
(((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^)
1 =
concat_p1 (wlx0 @ wry0) @
((((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^) @
(concat_p1 (wry1 @ wlx1))^) H_whiskerR_wlrnat_x_z := moveL_Mp
(wlrnat_x_z @
(concat_p1 (wrz1 @ wlx2))^)
(whiskerR wlrnat_x_z 1 )
(concat_p1 (wlx1 @ wrz0))
(moveL_pV
(concat_p1 (wrz1 @ wlx2))
((concat_p1 (wlx1 @ wrz0))^ @
whiskerR wlrnat_x_z 1 )
wlrnat_x_z
(whiskerR_p1 wlrnat_x_z)) : whiskerR wlrnat_x_z 1 =
concat_p1 (wlx1 @ wrz0) @
(wlrnat_x_z @
(concat_p1 (wrz1 @ wlx2))^)
whiskerR
((whiskerL wlx0 wrpp_yz0)^ @
((((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^ [I] wlrnat_x_z) @
whiskerR wrpp_yz1 wlx2)) 1 @
((whiskerR wrpp_yz1 1 )^ @
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 wlpp_yz1)
[-] ehlnat_x2) =
(ehlnat_x0
[-] (whiskerR wrpp_yz0 1 )^ @
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 wlpp_yz0)) @
whiskerL 1
((whiskerR wlpp_yz1 wrx2)^ @
((wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 wlpp_yz0))^
apply moveL_pV in wlrnat_V_x_z.X : Type a0, b0, c0, d0, e0, f0 : X wlx0 : a0 = b0 wlx1 : c0 = d0 wlx2 : e0 = f0 wrx0 : a0 = b0 wrx1 : c0 = d0 wrx2 : e0 = f0 wry0, wly0 : b0 = d0 wry1, wly1 : a0 = c0 wrz0, wlz0 : d0 = f0 wrz1, wlz1 : c0 = e0 wryz0, wlyz0 : b0 = f0 wryz1, wlyz1 : a0 = e0 ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : wry0 @ 1 = 1 @ wly0 ehrnat_y1 : wry1 @ 1 = 1 @ wly1 ehrnat_z0 : wrz0 @ 1 = 1 @ wlz0 ehrnat_z1 : wrz1 @ 1 = 1 @ wlz1 wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_x_z : wlx1 @ wrz0 = wrz1 @ wlx2 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0 wrpp_yz0 : wry0 @ wrz0 = wryz0 wlpp_yz0 : wly0 @ wlz0 = wlyz0 wrpp_yz1 : wry1 @ wrz1 = wryz1 wlpp_yz1 : wly1 @ wlz1 = wlyz1 wlrnat_V_x_z : whiskerR wlrnat_x_z 1 =
((ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^) @
(ehrnat_z1 [-] ehlnat_x2)^ H_whiskerR_wlrnat_x_y := moveL_Mp
((((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^) @
(concat_p1 (wry1 @ wlx1))^)
(whiskerR
(((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0
[-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1
[-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^)
1 )
(concat_p1 (wlx0 @ wry0))
(moveL_pV
(concat_p1 (wry1 @ wlx1))
((concat_p1 (wlx0 @ wry0))^ @
whiskerR
(((concat_p1
(wlx0 @ wry0))^ @
(((ehlnat_x0
[-] ehrnat_y0) @
whiskerL 1
wlrnat_y_x^) @
(ehrnat_y1
[-] ehlnat_x1)^)) @
((concat_p1
(wry1 @ wlx1))^)^)
1 )
(((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0
[-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1
[-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^)
(whiskerR_p1
(((concat_p1
(wlx0 @ wry0))^ @
(((ehlnat_x0
[-] ehrnat_y0) @
whiskerL 1
wlrnat_y_x^) @
(ehrnat_y1
[-] ehlnat_x1)^)) @
((concat_p1
(wry1 @ wlx1))^)^))) : whiskerR
(((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^)
1 =
concat_p1 (wlx0 @ wry0) @
((((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^) @
(concat_p1 (wry1 @ wlx1))^) H_whiskerR_wlrnat_x_z := moveL_Mp
(wlrnat_x_z @
(concat_p1 (wrz1 @ wlx2))^)
(whiskerR wlrnat_x_z 1 )
(concat_p1 (wlx1 @ wrz0))
(moveL_pV
(concat_p1 (wrz1 @ wlx2))
((concat_p1 (wlx1 @ wrz0))^ @
whiskerR wlrnat_x_z 1 )
wlrnat_x_z
(whiskerR_p1 wlrnat_x_z)) : whiskerR wlrnat_x_z 1 =
concat_p1 (wlx1 @ wrz0) @
(wlrnat_x_z @
(concat_p1 (wrz1 @ wlx2))^)
whiskerR
((whiskerL wlx0 wrpp_yz0)^ @
((((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^ [I] wlrnat_x_z) @
whiskerR wrpp_yz1 wlx2)) 1 @
((whiskerR wrpp_yz1 1 )^ @
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 wlpp_yz1)
[-] ehlnat_x2) =
(ehlnat_x0
[-] (whiskerR wrpp_yz0 1 )^ @
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 wlpp_yz0)) @
whiskerL 1
((whiskerR wlpp_yz1 wrx2)^ @
((wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 wlpp_yz0))^
apply (concat H_whiskerR_wlrnat_x_z^) in wlrnat_V_x_z.X : Type a0, b0, c0, d0, e0, f0 : X wlx0 : a0 = b0 wlx1 : c0 = d0 wlx2 : e0 = f0 wrx0 : a0 = b0 wrx1 : c0 = d0 wrx2 : e0 = f0 wry0, wly0 : b0 = d0 wry1, wly1 : a0 = c0 wrz0, wlz0 : d0 = f0 wrz1, wlz1 : c0 = e0 wryz0, wlyz0 : b0 = f0 wryz1, wlyz1 : a0 = e0 ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : wry0 @ 1 = 1 @ wly0 ehrnat_y1 : wry1 @ 1 = 1 @ wly1 ehrnat_z0 : wrz0 @ 1 = 1 @ wlz0 ehrnat_z1 : wrz1 @ 1 = 1 @ wlz1 wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_x_z : wlx1 @ wrz0 = wrz1 @ wlx2 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0 wrpp_yz0 : wry0 @ wrz0 = wryz0 wlpp_yz0 : wly0 @ wlz0 = wlyz0 wrpp_yz1 : wry1 @ wrz1 = wryz1 wlpp_yz1 : wly1 @ wlz1 = wlyz1 wlrnat_V_x_z : concat_p1 (wlx1 @ wrz0) @
(wlrnat_x_z @
(concat_p1 (wrz1 @ wlx2))^) =
((ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^) @
(ehrnat_z1 [-] ehlnat_x2)^ H_whiskerR_wlrnat_x_y := moveL_Mp
((((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^) @
(concat_p1 (wry1 @ wlx1))^)
(whiskerR
(((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0
[-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1
[-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^)
1 )
(concat_p1 (wlx0 @ wry0))
(moveL_pV
(concat_p1 (wry1 @ wlx1))
((concat_p1 (wlx0 @ wry0))^ @
whiskerR
(((concat_p1
(wlx0 @ wry0))^ @
(((ehlnat_x0
[-] ehrnat_y0) @
whiskerL 1
wlrnat_y_x^) @
(ehrnat_y1
[-] ehlnat_x1)^)) @
((concat_p1
(wry1 @ wlx1))^)^)
1 )
(((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0
[-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1
[-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^)
(whiskerR_p1
(((concat_p1
(wlx0 @ wry0))^ @
(((ehlnat_x0
[-] ehrnat_y0) @
whiskerL 1
wlrnat_y_x^) @
(ehrnat_y1
[-] ehlnat_x1)^)) @
((concat_p1
(wry1 @ wlx1))^)^))) : whiskerR
(((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^)
1 =
concat_p1 (wlx0 @ wry0) @
((((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^) @
(concat_p1 (wry1 @ wlx1))^) H_whiskerR_wlrnat_x_z := moveL_Mp
(wlrnat_x_z @
(concat_p1 (wrz1 @ wlx2))^)
(whiskerR wlrnat_x_z 1 )
(concat_p1 (wlx1 @ wrz0))
(moveL_pV
(concat_p1 (wrz1 @ wlx2))
((concat_p1 (wlx1 @ wrz0))^ @
whiskerR wlrnat_x_z 1 )
wlrnat_x_z
(whiskerR_p1 wlrnat_x_z)) : whiskerR wlrnat_x_z 1 =
concat_p1 (wlx1 @ wrz0) @
(wlrnat_x_z @
(concat_p1 (wrz1 @ wlx2))^)
whiskerR
((whiskerL wlx0 wrpp_yz0)^ @
((((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^ [I] wlrnat_x_z) @
whiskerR wrpp_yz1 wlx2)) 1 @
((whiskerR wrpp_yz1 1 )^ @
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 wlpp_yz1)
[-] ehlnat_x2) =
(ehlnat_x0
[-] (whiskerR wrpp_yz0 1 )^ @
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 wlpp_yz0)) @
whiskerL 1
((whiskerR wlpp_yz1 wrx2)^ @
((wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 wlpp_yz0))^
apply moveL_Vp, moveL_pV in wlrnat_V_x_z.X : Type a0, b0, c0, d0, e0, f0 : X wlx0 : a0 = b0 wlx1 : c0 = d0 wlx2 : e0 = f0 wrx0 : a0 = b0 wrx1 : c0 = d0 wrx2 : e0 = f0 wry0, wly0 : b0 = d0 wry1, wly1 : a0 = c0 wrz0, wlz0 : d0 = f0 wrz1, wlz1 : c0 = e0 wryz0, wlyz0 : b0 = f0 wryz1, wlyz1 : a0 = e0 ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : wry0 @ 1 = 1 @ wly0 ehrnat_y1 : wry1 @ 1 = 1 @ wly1 ehrnat_z0 : wrz0 @ 1 = 1 @ wlz0 ehrnat_z1 : wrz1 @ 1 = 1 @ wlz1 wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_x_z : wlx1 @ wrz0 = wrz1 @ wlx2 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0 wrpp_yz0 : wry0 @ wrz0 = wryz0 wlpp_yz0 : wly0 @ wlz0 = wlyz0 wrpp_yz1 : wry1 @ wrz1 = wryz1 wlpp_yz1 : wly1 @ wlz1 = wlyz1 wlrnat_V_x_z : wlrnat_x_z =
((concat_p1 (wlx1 @ wrz0))^ @
(((ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^) @
(ehrnat_z1 [-] ehlnat_x2)^)) @
((concat_p1 (wrz1 @ wlx2))^)^ H_whiskerR_wlrnat_x_y := moveL_Mp
((((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^) @
(concat_p1 (wry1 @ wlx1))^)
(whiskerR
(((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0
[-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1
[-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^)
1 )
(concat_p1 (wlx0 @ wry0))
(moveL_pV
(concat_p1 (wry1 @ wlx1))
((concat_p1 (wlx0 @ wry0))^ @
whiskerR
(((concat_p1
(wlx0 @ wry0))^ @
(((ehlnat_x0
[-] ehrnat_y0) @
whiskerL 1
wlrnat_y_x^) @
(ehrnat_y1
[-] ehlnat_x1)^)) @
((concat_p1
(wry1 @ wlx1))^)^)
1 )
(((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0
[-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1
[-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^)
(whiskerR_p1
(((concat_p1
(wlx0 @ wry0))^ @
(((ehlnat_x0
[-] ehrnat_y0) @
whiskerL 1
wlrnat_y_x^) @
(ehrnat_y1
[-] ehlnat_x1)^)) @
((concat_p1
(wry1 @ wlx1))^)^))) : whiskerR
(((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^)
1 =
concat_p1 (wlx0 @ wry0) @
((((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^) @
(concat_p1 (wry1 @ wlx1))^) H_whiskerR_wlrnat_x_z := moveL_Mp
(wlrnat_x_z @
(concat_p1 (wrz1 @ wlx2))^)
(whiskerR wlrnat_x_z 1 )
(concat_p1 (wlx1 @ wrz0))
(moveL_pV
(concat_p1 (wrz1 @ wlx2))
((concat_p1 (wlx1 @ wrz0))^ @
whiskerR wlrnat_x_z 1 )
wlrnat_x_z
(whiskerR_p1 wlrnat_x_z)) : whiskerR wlrnat_x_z 1 =
concat_p1 (wlx1 @ wrz0) @
(wlrnat_x_z @
(concat_p1 (wrz1 @ wlx2))^)
whiskerR
((whiskerL wlx0 wrpp_yz0)^ @
((((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^ [I] wlrnat_x_z) @
whiskerR wrpp_yz1 wlx2)) 1 @
((whiskerR wrpp_yz1 1 )^ @
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 wlpp_yz1)
[-] ehlnat_x2) =
(ehlnat_x0
[-] (whiskerR wrpp_yz0 1 )^ @
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 wlpp_yz0)) @
whiskerL 1
((whiskerR wlpp_yz1 wrx2)^ @
((wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 wlpp_yz0))^
apply symmetry in wlrnat_V_x_z.X : Type a0, b0, c0, d0, e0, f0 : X wlx0 : a0 = b0 wlx1 : c0 = d0 wlx2 : e0 = f0 wrx0 : a0 = b0 wrx1 : c0 = d0 wrx2 : e0 = f0 wry0, wly0 : b0 = d0 wry1, wly1 : a0 = c0 wrz0, wlz0 : d0 = f0 wrz1, wlz1 : c0 = e0 wryz0, wlyz0 : b0 = f0 wryz1, wlyz1 : a0 = e0 ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : wry0 @ 1 = 1 @ wly0 ehrnat_y1 : wry1 @ 1 = 1 @ wly1 ehrnat_z0 : wrz0 @ 1 = 1 @ wlz0 ehrnat_z1 : wrz1 @ 1 = 1 @ wlz1 wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_x_z : wlx1 @ wrz0 = wrz1 @ wlx2 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0 wrpp_yz0 : wry0 @ wrz0 = wryz0 wlpp_yz0 : wly0 @ wlz0 = wlyz0 wrpp_yz1 : wry1 @ wrz1 = wryz1 wlpp_yz1 : wly1 @ wlz1 = wlyz1 wlrnat_V_x_z : ((concat_p1 (wlx1 @ wrz0))^ @
(((ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^) @
(ehrnat_z1 [-] ehlnat_x2)^)) @
((concat_p1 (wrz1 @ wlx2))^)^ =
wlrnat_x_z H_whiskerR_wlrnat_x_y := moveL_Mp
((((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^) @
(concat_p1 (wry1 @ wlx1))^)
(whiskerR
(((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0
[-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1
[-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^)
1 )
(concat_p1 (wlx0 @ wry0))
(moveL_pV
(concat_p1 (wry1 @ wlx1))
((concat_p1 (wlx0 @ wry0))^ @
whiskerR
(((concat_p1
(wlx0 @ wry0))^ @
(((ehlnat_x0
[-] ehrnat_y0) @
whiskerL 1
wlrnat_y_x^) @
(ehrnat_y1
[-] ehlnat_x1)^)) @
((concat_p1
(wry1 @ wlx1))^)^)
1 )
(((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0
[-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1
[-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^)
(whiskerR_p1
(((concat_p1
(wlx0 @ wry0))^ @
(((ehlnat_x0
[-] ehrnat_y0) @
whiskerL 1
wlrnat_y_x^) @
(ehrnat_y1
[-] ehlnat_x1)^)) @
((concat_p1
(wry1 @ wlx1))^)^))) : whiskerR
(((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^)
1 =
concat_p1 (wlx0 @ wry0) @
((((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^) @
(concat_p1 (wry1 @ wlx1))^) H_whiskerR_wlrnat_x_z := moveL_Mp
(wlrnat_x_z @
(concat_p1 (wrz1 @ wlx2))^)
(whiskerR wlrnat_x_z 1 )
(concat_p1 (wlx1 @ wrz0))
(moveL_pV
(concat_p1 (wrz1 @ wlx2))
((concat_p1 (wlx1 @ wrz0))^ @
whiskerR wlrnat_x_z 1 )
wlrnat_x_z
(whiskerR_p1 wlrnat_x_z)) : whiskerR wlrnat_x_z 1 =
concat_p1 (wlx1 @ wrz0) @
(wlrnat_x_z @
(concat_p1 (wrz1 @ wlx2))^)
whiskerR
((whiskerL wlx0 wrpp_yz0)^ @
((((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^ [I] wlrnat_x_z) @
whiskerR wrpp_yz1 wlx2)) 1 @
((whiskerR wrpp_yz1 1 )^ @
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 wlpp_yz1)
[-] ehlnat_x2) =
(ehlnat_x0
[-] (whiskerR wrpp_yz0 1 )^ @
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 wlpp_yz0)) @
whiskerL 1
((whiskerR wlpp_yz1 wrx2)^ @
((wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 wlpp_yz0))^
destruct wlrnat_V_x_z.X : Type a0, b0, c0, d0, e0, f0 : X wlx0 : a0 = b0 wlx1 : c0 = d0 wlx2 : e0 = f0 wrx0 : a0 = b0 wrx1 : c0 = d0 wrx2 : e0 = f0 wry0, wly0 : b0 = d0 wry1, wly1 : a0 = c0 wrz0, wlz0 : d0 = f0 wrz1, wlz1 : c0 = e0 wryz0, wlyz0 : b0 = f0 wryz1, wlyz1 : a0 = e0 ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : wry0 @ 1 = 1 @ wly0 ehrnat_y1 : wry1 @ 1 = 1 @ wly1 ehrnat_z0 : wrz0 @ 1 = 1 @ wlz0 ehrnat_z1 : wrz1 @ 1 = 1 @ wlz1 wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0 wrpp_yz0 : wry0 @ wrz0 = wryz0 wlpp_yz0 : wly0 @ wlz0 = wlyz0 wrpp_yz1 : wry1 @ wrz1 = wryz1 wlpp_yz1 : wly1 @ wlz1 = wlyz1 wlrnat_V_x_z : ((concat_p1 (wlx1 @ wrz0))^ @
(((ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^) @
(ehrnat_z1 [-] ehlnat_x2)^)) @
((concat_p1 (wrz1 @ wlx2))^)^ =
((concat_p1 (wlx1 @ wrz0))^ @
(((ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^) @
(ehrnat_z1 [-] ehlnat_x2)^)) @
((concat_p1 (wrz1 @ wlx2))^)^ H_whiskerR_wlrnat_x_y := moveL_Mp
((((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^) @
(concat_p1 (wry1 @ wlx1))^)
(whiskerR
(((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0
[-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1
[-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^)
1 )
(concat_p1 (wlx0 @ wry0))
(moveL_pV
(concat_p1 (wry1 @ wlx1))
((concat_p1 (wlx0 @ wry0))^ @
whiskerR
(((concat_p1
(wlx0 @ wry0))^ @
(((ehlnat_x0
[-] ehrnat_y0) @
whiskerL 1
wlrnat_y_x^) @
(ehrnat_y1
[-] ehlnat_x1)^)) @
((concat_p1
(wry1 @ wlx1))^)^)
1 )
(((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0
[-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1
[-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^)
(whiskerR_p1
(((concat_p1
(wlx0 @ wry0))^ @
(((ehlnat_x0
[-] ehrnat_y0) @
whiskerL 1
wlrnat_y_x^) @
(ehrnat_y1
[-] ehlnat_x1)^)) @
((concat_p1
(wry1 @ wlx1))^)^))) : whiskerR
(((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^)
1 =
concat_p1 (wlx0 @ wry0) @
((((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^) @
(concat_p1 (wry1 @ wlx1))^) H_whiskerR_wlrnat_x_z := moveL_Mp
((((concat_p1 (wlx1 @ wrz0))^ @
(((ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^) @
(ehrnat_z1 [-] ehlnat_x2)^)) @
((concat_p1 (wrz1 @ wlx2))^)^) @
(concat_p1 (wrz1 @ wlx2))^)
(whiskerR
(((concat_p1 (wlx1 @ wrz0))^ @
(((ehlnat_x1
[-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^) @
(ehrnat_z1
[-] ehlnat_x2)^)) @
((concat_p1 (wrz1 @ wlx2))^)^)
1 )
(concat_p1 (wlx1 @ wrz0))
(moveL_pV
(concat_p1 (wrz1 @ wlx2))
((concat_p1 (wlx1 @ wrz0))^ @
whiskerR
(((concat_p1
(wlx1 @ wrz0))^ @
(((ehlnat_x1
[-] ehrnat_z0) @
whiskerL 1
wlrnat_z_x^) @
(ehrnat_z1
[-] ehlnat_x2)^)) @
((concat_p1
(wrz1 @ wlx2))^)^)
1 )
(((concat_p1 (wlx1 @ wrz0))^ @
(((ehlnat_x1
[-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^) @
(ehrnat_z1
[-] ehlnat_x2)^)) @
((concat_p1 (wrz1 @ wlx2))^)^)
(whiskerR_p1
(((concat_p1
(wlx1 @ wrz0))^ @
(((ehlnat_x1
[-] ehrnat_z0) @
whiskerL 1
wlrnat_z_x^) @
(ehrnat_z1
[-] ehlnat_x2)^)) @
((concat_p1
(wrz1 @ wlx2))^)^))) : whiskerR
(((concat_p1 (wlx1 @ wrz0))^ @
(((ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^) @
(ehrnat_z1 [-] ehlnat_x2)^)) @
((concat_p1 (wrz1 @ wlx2))^)^)
1 =
concat_p1 (wlx1 @ wrz0) @
((((concat_p1 (wlx1 @ wrz0))^ @
(((ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^) @
(ehrnat_z1 [-] ehlnat_x2)^)) @
((concat_p1 (wrz1 @ wlx2))^)^) @
(concat_p1 (wrz1 @ wlx2))^)
whiskerR
((whiskerL wlx0 wrpp_yz0)^ @
((((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^
[I] ((concat_p1 (wlx1 @ wrz0))^ @
(((ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^) @
(ehrnat_z1 [-] ehlnat_x2)^)) @
((concat_p1 (wrz1 @ wlx2))^)^) @
whiskerR wrpp_yz1 wlx2)) 1 @
((whiskerR wrpp_yz1 1 )^ @
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 wlpp_yz1)
[-] ehlnat_x2) =
(ehlnat_x0
[-] (whiskerR wrpp_yz0 1 )^ @
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 wlpp_yz0)) @
whiskerL 1
((whiskerR wlpp_yz1 wrx2)^ @
((wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 wlpp_yz0))^
clear wlrnat_V_x_z.X : Type a0, b0, c0, d0, e0, f0 : X wlx0 : a0 = b0 wlx1 : c0 = d0 wlx2 : e0 = f0 wrx0 : a0 = b0 wrx1 : c0 = d0 wrx2 : e0 = f0 wry0, wly0 : b0 = d0 wry1, wly1 : a0 = c0 wrz0, wlz0 : d0 = f0 wrz1, wlz1 : c0 = e0 wryz0, wlyz0 : b0 = f0 wryz1, wlyz1 : a0 = e0 ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : wry0 @ 1 = 1 @ wly0 ehrnat_y1 : wry1 @ 1 = 1 @ wly1 ehrnat_z0 : wrz0 @ 1 = 1 @ wlz0 ehrnat_z1 : wrz1 @ 1 = 1 @ wlz1 wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0 wrpp_yz0 : wry0 @ wrz0 = wryz0 wlpp_yz0 : wly0 @ wlz0 = wlyz0 wrpp_yz1 : wry1 @ wrz1 = wryz1 wlpp_yz1 : wly1 @ wlz1 = wlyz1 H_whiskerR_wlrnat_x_y := moveL_Mp
((((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^) @
(concat_p1 (wry1 @ wlx1))^)
(whiskerR
(((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0
[-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1
[-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^)
1 )
(concat_p1 (wlx0 @ wry0))
(moveL_pV
(concat_p1 (wry1 @ wlx1))
((concat_p1 (wlx0 @ wry0))^ @
whiskerR
(((concat_p1
(wlx0 @ wry0))^ @
(((ehlnat_x0
[-] ehrnat_y0) @
whiskerL 1
wlrnat_y_x^) @
(ehrnat_y1
[-] ehlnat_x1)^)) @
((concat_p1
(wry1 @ wlx1))^)^)
1 )
(((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0
[-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1
[-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^)
(whiskerR_p1
(((concat_p1
(wlx0 @ wry0))^ @
(((ehlnat_x0
[-] ehrnat_y0) @
whiskerL 1
wlrnat_y_x^) @
(ehrnat_y1
[-] ehlnat_x1)^)) @
((concat_p1
(wry1 @ wlx1))^)^))) : whiskerR
(((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^)
1 =
concat_p1 (wlx0 @ wry0) @
((((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^) @
(concat_p1 (wry1 @ wlx1))^) H_whiskerR_wlrnat_x_z := moveL_Mp
((((concat_p1 (wlx1 @ wrz0))^ @
(((ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^) @
(ehrnat_z1 [-] ehlnat_x2)^)) @
((concat_p1 (wrz1 @ wlx2))^)^) @
(concat_p1 (wrz1 @ wlx2))^)
(whiskerR
(((concat_p1 (wlx1 @ wrz0))^ @
(((ehlnat_x1
[-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^) @
(ehrnat_z1
[-] ehlnat_x2)^)) @
((concat_p1 (wrz1 @ wlx2))^)^)
1 )
(concat_p1 (wlx1 @ wrz0))
(moveL_pV
(concat_p1 (wrz1 @ wlx2))
((concat_p1 (wlx1 @ wrz0))^ @
whiskerR
(((concat_p1
(wlx1 @ wrz0))^ @
(((ehlnat_x1
[-] ehrnat_z0) @
whiskerL 1
wlrnat_z_x^) @
(ehrnat_z1
[-] ehlnat_x2)^)) @
((concat_p1
(wrz1 @ wlx2))^)^)
1 )
(((concat_p1 (wlx1 @ wrz0))^ @
(((ehlnat_x1
[-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^) @
(ehrnat_z1
[-] ehlnat_x2)^)) @
((concat_p1 (wrz1 @ wlx2))^)^)
(whiskerR_p1
(((concat_p1
(wlx1 @ wrz0))^ @
(((ehlnat_x1
[-] ehrnat_z0) @
whiskerL 1
wlrnat_z_x^) @
(ehrnat_z1
[-] ehlnat_x2)^)) @
((concat_p1
(wrz1 @ wlx2))^)^))) : whiskerR
(((concat_p1 (wlx1 @ wrz0))^ @
(((ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^) @
(ehrnat_z1 [-] ehlnat_x2)^)) @
((concat_p1 (wrz1 @ wlx2))^)^)
1 =
concat_p1 (wlx1 @ wrz0) @
((((concat_p1 (wlx1 @ wrz0))^ @
(((ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^) @
(ehrnat_z1 [-] ehlnat_x2)^)) @
((concat_p1 (wrz1 @ wlx2))^)^) @
(concat_p1 (wrz1 @ wlx2))^)
whiskerR
((whiskerL wlx0 wrpp_yz0)^ @
((((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^
[I] ((concat_p1 (wlx1 @ wrz0))^ @
(((ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^) @
(ehrnat_z1 [-] ehlnat_x2)^)) @
((concat_p1 (wrz1 @ wlx2))^)^) @
whiskerR wrpp_yz1 wlx2)) 1 @
((whiskerR wrpp_yz1 1 )^ @
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 wlpp_yz1)
[-] ehlnat_x2) =
(ehlnat_x0
[-] (whiskerR wrpp_yz0 1 )^ @
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 wlpp_yz0)) @
whiskerL 1
((whiskerR wlpp_yz1 wrx2)^ @
((wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 wlpp_yz0))^
clear H_whiskerR_wlrnat_x_y H_whiskerR_wlrnat_x_z.X : Type a0, b0, c0, d0, e0, f0 : X wlx0 : a0 = b0 wlx1 : c0 = d0 wlx2 : e0 = f0 wrx0 : a0 = b0 wrx1 : c0 = d0 wrx2 : e0 = f0 wry0, wly0 : b0 = d0 wry1, wly1 : a0 = c0 wrz0, wlz0 : d0 = f0 wrz1, wlz1 : c0 = e0 wryz0, wlyz0 : b0 = f0 wryz1, wlyz1 : a0 = e0 ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : wry0 @ 1 = 1 @ wly0 ehrnat_y1 : wry1 @ 1 = 1 @ wly1 ehrnat_z0 : wrz0 @ 1 = 1 @ wlz0 ehrnat_z1 : wrz1 @ 1 = 1 @ wlz1 wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0 wrpp_yz0 : wry0 @ wrz0 = wryz0 wlpp_yz0 : wly0 @ wlz0 = wlyz0 wrpp_yz1 : wry1 @ wrz1 = wryz1 wlpp_yz1 : wly1 @ wlz1 = wlyz1
whiskerR
((whiskerL wlx0 wrpp_yz0)^ @
((((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^
[I] ((concat_p1 (wlx1 @ wrz0))^ @
(((ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^) @
(ehrnat_z1 [-] ehlnat_x2)^)) @
((concat_p1 (wrz1 @ wlx2))^)^) @
whiskerR wrpp_yz1 wlx2)) 1 @
((whiskerR wrpp_yz1 1 )^ @
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 wlpp_yz1)
[-] ehlnat_x2) =
(ehlnat_x0
[-] (whiskerR wrpp_yz0 1 )^ @
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 wlpp_yz0)) @
whiskerL 1
((whiskerR wlpp_yz1 wrx2)^ @
((wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 wlpp_yz0))^
destruct wrpp_yz0, wlpp_yz0, wrpp_yz1, wlpp_yz1.X : Type a0, b0, c0, d0, e0, f0 : X wlx0 : a0 = b0 wlx1 : c0 = d0 wlx2 : e0 = f0 wrx0 : a0 = b0 wrx1 : c0 = d0 wrx2 : e0 = f0 wry0, wly0 : b0 = d0 wry1, wly1 : a0 = c0 wrz0, wlz0 : d0 = f0 wrz1, wlz1 : c0 = e0 ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : wry0 @ 1 = 1 @ wly0 ehrnat_y1 : wry1 @ 1 = 1 @ wly1 ehrnat_z0 : wrz0 @ 1 = 1 @ wlz0 ehrnat_z1 : wrz1 @ 1 = 1 @ wlz1 wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0 wrpp_yz0 : wry0 @ wrz0 = wry0 @ wrz0 wlpp_yz0 : wly0 @ wlz0 = wly0 @ wlz0 wrpp_yz1 : wry1 @ wrz1 = wry1 @ wrz1 wlpp_yz1 : wly1 @ wlz1 = wly1 @ wlz1
whiskerR
((whiskerL wlx0 1 )^ @
((((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^
[I] ((concat_p1 (wlx1 @ wrz0))^ @
(((ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^) @
(ehrnat_z1 [-] ehlnat_x2)^)) @
((concat_p1 (wrz1 @ wlx2))^)^) @
whiskerR 1 wlx2)) 1 @
((whiskerR 1 1 )^ @
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 1 )
[-] ehlnat_x2) =
(ehlnat_x0
[-] (whiskerR 1 1 )^ @
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 1 )) @
whiskerL 1
((whiskerR 1 wrx2)^ @
((wlrnat_y_x [-] wlrnat_z_x) @ whiskerL wrx0 1 ))^
clear wrpp_yz0 wlpp_yz0 wrpp_yz1 wlpp_yz1.X : Type a0, b0, c0, d0, e0, f0 : X wlx0 : a0 = b0 wlx1 : c0 = d0 wlx2 : e0 = f0 wrx0 : a0 = b0 wrx1 : c0 = d0 wrx2 : e0 = f0 wry0, wly0 : b0 = d0 wry1, wly1 : a0 = c0 wrz0, wlz0 : d0 = f0 wrz1, wlz1 : c0 = e0 ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : wry0 @ 1 = 1 @ wly0 ehrnat_y1 : wry1 @ 1 = 1 @ wly1 ehrnat_z0 : wrz0 @ 1 = 1 @ wlz0 ehrnat_z1 : wrz1 @ 1 = 1 @ wlz1 wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0
whiskerR
((whiskerL wlx0 1 )^ @
((((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^
[I] ((concat_p1 (wlx1 @ wrz0))^ @
(((ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^) @
(ehrnat_z1 [-] ehlnat_x2)^)) @
((concat_p1 (wrz1 @ wlx2))^)^) @
whiskerR 1 wlx2)) 1 @
((whiskerR 1 1 )^ @
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 1 )
[-] ehlnat_x2) =
(ehlnat_x0
[-] (whiskerR 1 1 )^ @
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 1 )) @
whiskerL 1
((whiskerR 1 wrx2)^ @
((wlrnat_y_x [-] wlrnat_z_x) @ whiskerL wrx0 1 ))^
revert wlrnat_y_x wlrnat_z_x.X : Type a0, b0, c0, d0, e0, f0 : X wlx0 : a0 = b0 wlx1 : c0 = d0 wlx2 : e0 = f0 wrx0 : a0 = b0 wrx1 : c0 = d0 wrx2 : e0 = f0 wry0, wly0 : b0 = d0 wry1, wly1 : a0 = c0 wrz0, wlz0 : d0 = f0 wrz1, wlz1 : c0 = e0 ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : wry0 @ 1 = 1 @ wly0 ehrnat_y1 : wry1 @ 1 = 1 @ wly1 ehrnat_z0 : wrz0 @ 1 = 1 @ wlz0 ehrnat_z1 : wrz1 @ 1 = 1 @ wlz1
forall (wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0)
(wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0),
whiskerR
((whiskerL wlx0 1 )^ @
((((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^
[I] ((concat_p1 (wlx1 @ wrz0))^ @
(((ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^) @
(ehrnat_z1 [-] ehlnat_x2)^)) @
((concat_p1 (wrz1 @ wlx2))^)^) @
whiskerR 1 wlx2)) 1 @
((whiskerR 1 1 )^ @
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 1 )
[-] ehlnat_x2) =
(ehlnat_x0
[-] (whiskerR 1 1 )^ @
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 1 )) @
whiskerL 1
((whiskerR 1 wrx2)^ @
((wlrnat_y_x [-] wlrnat_z_x) @ whiskerL wrx0 1 ))^
revert wrx0 ehlnat_x0.X : Type a0, b0, c0, d0, e0, f0 : X wlx0 : a0 = b0 wlx1 : c0 = d0 wlx2 : e0 = f0 wrx1 : c0 = d0 wrx2 : e0 = f0 wry0, wly0 : b0 = d0 wry1, wly1 : a0 = c0 wrz0, wlz0 : d0 = f0 wrz1, wlz1 : c0 = e0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : wry0 @ 1 = 1 @ wly0 ehrnat_y1 : wry1 @ 1 = 1 @ wly1 ehrnat_z0 : wrz0 @ 1 = 1 @ wlz0 ehrnat_z1 : wrz1 @ 1 = 1 @ wlz1
forall (wrx0 : a0 = b0)
(ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0)
(wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0)
(wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0),
whiskerR
((whiskerL wlx0 1 )^ @
((((concat_p1 (wlx0 @ wry0))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^
[I] ((concat_p1 (wlx1 @ wrz0))^ @
(((ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^) @
(ehrnat_z1 [-] ehlnat_x2)^)) @
((concat_p1 (wrz1 @ wlx2))^)^) @
whiskerR 1 wlx2)) 1 @
((whiskerR 1 1 )^ @
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 1 )
[-] ehlnat_x2) =
(ehlnat_x0
[-] (whiskerR 1 1 )^ @
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 1 )) @
whiskerL 1
((whiskerR 1 wrx2)^ @
((wlrnat_y_x [-] wlrnat_z_x) @ whiskerL wrx0 1 ))^
snapply equiv_path_ind_rlucancel. X : Type a0, b0, c0, d0, e0, f0 : X wlx0 : a0 = b0 wlx1 : c0 = d0 wlx2 : e0 = f0 wrx1 : c0 = d0 wrx2 : e0 = f0 wry0, wly0 : b0 = d0 wry1, wly1 : a0 = c0 wrz0, wlz0 : d0 = f0 wrz1, wlz1 : c0 = e0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : wry0 @ 1 = 1 @ wly0 ehrnat_y1 : wry1 @ 1 = 1 @ wly1 ehrnat_z0 : wrz0 @ 1 = 1 @ wlz0 ehrnat_z1 : wrz1 @ 1 = 1 @ wlz1
(fun (q : a0 = b0) (s : wlx0 @ 1 = 1 @ q) =>
forall (wlrnat_y_x : wly1 @ wrx1 = q @ wly0)
(wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0),
whiskerR
((whiskerL wlx0 1 )^ @
((((concat_p1 (wlx0 @ wry0))^ @
(((s [-] ehrnat_y0) @ whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^
[I] ((concat_p1 (wlx1 @ wrz0))^ @
(((ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^) @
(ehrnat_z1 [-] ehlnat_x2)^)) @
((concat_p1 (wrz1 @ wlx2))^)^) @
whiskerR 1 wlx2)) 1 @
((whiskerR 1 1 )^ @
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 1 )
[-] ehlnat_x2) =
(s
[-] (whiskerR 1 1 )^ @
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 1 )) @
whiskerL 1
((whiskerR 1 wrx2)^ @
((wlrnat_y_x [-] wlrnat_z_x) @ whiskerL q 1 ))^)
wlx0 (rlucancel 1 )
revert wrx1 ehlnat_x1.X : Type a0, b0, c0, d0, e0, f0 : X wlx0 : a0 = b0 wlx1 : c0 = d0 wlx2, wrx2 : e0 = f0 wry0, wly0 : b0 = d0 wry1, wly1 : a0 = c0 wrz0, wlz0 : d0 = f0 wrz1, wlz1 : c0 = e0 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : wry0 @ 1 = 1 @ wly0 ehrnat_y1 : wry1 @ 1 = 1 @ wly1 ehrnat_z0 : wrz0 @ 1 = 1 @ wlz0 ehrnat_z1 : wrz1 @ 1 = 1 @ wlz1
forall (wrx1 : c0 = d0)
(ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1),
(fun (q : a0 = b0) (s : wlx0 @ 1 = 1 @ q) =>
forall (wlrnat_y_x : wly1 @ wrx1 = q @ wly0)
(wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0),
whiskerR
((whiskerL wlx0 1 )^ @
((((concat_p1 (wlx0 @ wry0))^ @
(((s [-] ehrnat_y0) @ whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
((concat_p1 (wry1 @ wlx1))^)^
[I] ((concat_p1 (wlx1 @ wrz0))^ @
(((ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^) @
(ehrnat_z1 [-] ehlnat_x2)^)) @
((concat_p1 (wrz1 @ wlx2))^)^) @
whiskerR 1 wlx2)) 1 @
((whiskerR 1 1 )^ @
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 1 )
[-] ehlnat_x2) =
(s
[-] (whiskerR 1 1 )^ @
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 1 )) @
whiskerL 1
((whiskerR 1 wrx2)^ @
((wlrnat_y_x [-] wlrnat_z_x) @ whiskerL q 1 ))^)
wlx0 (rlucancel 1 )
snapply equiv_path_ind_rlucancel. X : Type a0, b0, c0, d0, e0, f0 : X wlx0 : a0 = b0 wlx1 : c0 = d0 wlx2, wrx2 : e0 = f0 wry0, wly0 : b0 = d0 wry1, wly1 : a0 = c0 wrz0, wlz0 : d0 = f0 wrz1, wlz1 : c0 = e0 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : wry0 @ 1 = 1 @ wly0 ehrnat_y1 : wry1 @ 1 = 1 @ wly1 ehrnat_z0 : wrz0 @ 1 = 1 @ wlz0 ehrnat_z1 : wrz1 @ 1 = 1 @ wlz1
(fun (q : c0 = d0) (s : wlx1 @ 1 = 1 @ q) =>
(fun (q0 : a0 = b0) (s0 : wlx0 @ 1 = 1 @ q0) =>
forall (wlrnat_y_x : wly1 @ q = q0 @ wly0)
(wlrnat_z_x : wlz1 @ wrx2 = q @ wlz0),
whiskerR
((whiskerL wlx0 1 )^ @
((((concat_p1 (wlx0 @ wry0))^ @
(((s0 [-] ehrnat_y0) @ whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] s)^)) @
((concat_p1 (wry1 @ wlx1))^)^
[I] ((concat_p1 (wlx1 @ wrz0))^ @
(((s [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^) @
(ehrnat_z1 [-] ehlnat_x2)^)) @
((concat_p1 (wrz1 @ wlx2))^)^) @
whiskerR 1 wlx2)) 1 @
((whiskerR 1 1 )^ @
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 1 )
[-] ehlnat_x2) =
(s0
[-] (whiskerR 1 1 )^ @
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 1 )) @
whiskerL 1
((whiskerR 1 wrx2)^ @
((wlrnat_y_x [-] wlrnat_z_x) @ whiskerL q0 1 ))^)
wlx0 (rlucancel 1 )) wlx1 (rlucancel 1 )
revert wrx2 ehlnat_x2.X : Type a0, b0, c0, d0, e0, f0 : X wlx0 : a0 = b0 wlx1 : c0 = d0 wlx2 : e0 = f0 wry0, wly0 : b0 = d0 wry1, wly1 : a0 = c0 wrz0, wlz0 : d0 = f0 wrz1, wlz1 : c0 = e0 ehrnat_y0 : wry0 @ 1 = 1 @ wly0 ehrnat_y1 : wry1 @ 1 = 1 @ wly1 ehrnat_z0 : wrz0 @ 1 = 1 @ wlz0 ehrnat_z1 : wrz1 @ 1 = 1 @ wlz1
forall (wrx2 : e0 = f0)
(ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2),
(fun (q : c0 = d0) (s : wlx1 @ 1 = 1 @ q) =>
(fun (q0 : a0 = b0) (s0 : wlx0 @ 1 = 1 @ q0) =>
forall (wlrnat_y_x : wly1 @ q = q0 @ wly0)
(wlrnat_z_x : wlz1 @ wrx2 = q @ wlz0),
whiskerR
((whiskerL wlx0 1 )^ @
((((concat_p1 (wlx0 @ wry0))^ @
(((s0 [-] ehrnat_y0) @ whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] s)^)) @
((concat_p1 (wry1 @ wlx1))^)^
[I] ((concat_p1 (wlx1 @ wrz0))^ @
(((s [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^) @
(ehrnat_z1 [-] ehlnat_x2)^)) @
((concat_p1 (wrz1 @ wlx2))^)^) @
whiskerR 1 wlx2)) 1 @
((whiskerR 1 1 )^ @
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 1 )
[-] ehlnat_x2) =
(s0
[-] (whiskerR 1 1 )^ @
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 1 )) @
whiskerL 1
((whiskerR 1 wrx2)^ @
((wlrnat_y_x [-] wlrnat_z_x) @ whiskerL q0 1 ))^)
wlx0 (rlucancel 1 )) wlx1 (rlucancel 1 )
snapply equiv_path_ind_rlucancel. X : Type a0, b0, c0, d0, e0, f0 : X wlx0 : a0 = b0 wlx1 : c0 = d0 wlx2 : e0 = f0 wry0, wly0 : b0 = d0 wry1, wly1 : a0 = c0 wrz0, wlz0 : d0 = f0 wrz1, wlz1 : c0 = e0 ehrnat_y0 : wry0 @ 1 = 1 @ wly0 ehrnat_y1 : wry1 @ 1 = 1 @ wly1 ehrnat_z0 : wrz0 @ 1 = 1 @ wlz0 ehrnat_z1 : wrz1 @ 1 = 1 @ wlz1
(fun (q : e0 = f0) (s : wlx2 @ 1 = 1 @ q) =>
(fun (q0 : c0 = d0) (s0 : wlx1 @ 1 = 1 @ q0) =>
(fun (q1 : a0 = b0) (s1 : wlx0 @ 1 = 1 @ q1) =>
forall (wlrnat_y_x : wly1 @ q0 = q1 @ wly0)
(wlrnat_z_x : wlz1 @ q = q0 @ wlz0),
whiskerR
((whiskerL wlx0 1 )^ @
((((concat_p1 (wlx0 @ wry0))^ @
(((...) @ whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] s0)^)) @
((concat_p1 (wry1 @ wlx1))^)^
[I] ((concat_p1 (wlx1 @ wrz0))^ @
(((...) @ whiskerL 1 wlrnat_z_x^) @
(ehrnat_z1 [-] s)^)) @
((concat_p1 (wrz1 @ wlx2))^)^) @
whiskerR 1 wlx2)) 1 @
((whiskerR 1 1 )^ @
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 1 ) [-] s) =
(s1
[-] (whiskerR 1 1 )^ @
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 1 )) @
whiskerL 1
((whiskerR 1 q)^ @
((wlrnat_y_x [-] wlrnat_z_x) @ whiskerL q1 1 ))^)
wlx0 (rlucancel 1 )) wlx1 (rlucancel 1 )) wlx2
(rlucancel 1 )
revert wly0 ehrnat_y0.X : Type a0, b0, c0, d0, e0, f0 : X wlx0 : a0 = b0 wlx1 : c0 = d0 wlx2 : e0 = f0 wry0 : b0 = d0 wry1, wly1 : a0 = c0 wrz0, wlz0 : d0 = f0 wrz1, wlz1 : c0 = e0 ehrnat_y1 : wry1 @ 1 = 1 @ wly1 ehrnat_z0 : wrz0 @ 1 = 1 @ wlz0 ehrnat_z1 : wrz1 @ 1 = 1 @ wlz1
forall (wly0 : b0 = d0)
(ehrnat_y0 : wry0 @ 1 = 1 @ wly0),
(fun (q : e0 = f0) (s : wlx2 @ 1 = 1 @ q) =>
(fun (q0 : c0 = d0) (s0 : wlx1 @ 1 = 1 @ q0) =>
(fun (q1 : a0 = b0) (s1 : wlx0 @ 1 = 1 @ q1) =>
forall (wlrnat_y_x : wly1 @ q0 = q1 @ wly0)
(wlrnat_z_x : wlz1 @ q = q0 @ wlz0),
whiskerR
((whiskerL wlx0 1 )^ @
((((concat_p1 (wlx0 @ wry0))^ @
((... @ ...) @ (ehrnat_y1 [-] s0)^)) @
((concat_p1 (wry1 @ wlx1))^)^
[I] ((concat_p1 (wlx1 @ wrz0))^ @
((... @ ...) @ (ehrnat_z1 [-] s)^)) @
((concat_p1 (wrz1 @ wlx2))^)^) @
whiskerR 1 wlx2)) 1 @
((whiskerR 1 1 )^ @
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 1 ) [-] s) =
(s1
[-] (whiskerR 1 1 )^ @
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 1 )) @
whiskerL 1
((whiskerR 1 q)^ @
((wlrnat_y_x [-] wlrnat_z_x) @ whiskerL q1 1 ))^)
wlx0 (rlucancel 1 )) wlx1 (rlucancel 1 )) wlx2
(rlucancel 1 )
snapply equiv_path_ind_rlucancel. X : Type a0, b0, c0, d0, e0, f0 : X wlx0 : a0 = b0 wlx1 : c0 = d0 wlx2 : e0 = f0 wry0 : b0 = d0 wry1, wly1 : a0 = c0 wrz0, wlz0 : d0 = f0 wrz1, wlz1 : c0 = e0 ehrnat_y1 : wry1 @ 1 = 1 @ wly1 ehrnat_z0 : wrz0 @ 1 = 1 @ wlz0 ehrnat_z1 : wrz1 @ 1 = 1 @ wlz1
(fun (q : b0 = d0) (s : wry0 @ 1 = 1 @ q) =>
(fun (q0 : e0 = f0) (s0 : wlx2 @ 1 = 1 @ q0) =>
(fun (q1 : c0 = d0) (s1 : wlx1 @ 1 = 1 @ q1) =>
(fun (q2 : a0 = b0) (s2 : wlx0 @ 1 = 1 @ q2) =>
forall (wlrnat_y_x : wly1 @ q1 = q2 @ q)
(wlrnat_z_x : wlz1 @ q0 = q1 @ wlz0),
whiskerR
((whiskerL wlx0 1 )^ @
((((concat_p1 ...)^ @ (... @ ...^)) @
((concat_p1 (...))^)^
[I] ((concat_p1 ...)^ @ (... @ ...^)) @
((concat_p1 (...))^)^) @ whiskerR 1 wlx2))
1 @
((whiskerR 1 1 )^ @
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 1 ) [-] s0) =
(s2
[-] (whiskerR 1 1 )^ @
((s [-] ehrnat_z0) @ whiskerL 1 1 )) @
whiskerL 1
((whiskerR 1 q0)^ @
((wlrnat_y_x [-] wlrnat_z_x) @ whiskerL q2 1 ))^)
wlx0 (rlucancel 1 )) wlx1 (rlucancel 1 )) wlx2
(rlucancel 1 )) wry0 (rlucancel 1 )
revert wly1 ehrnat_y1.X : Type a0, b0, c0, d0, e0, f0 : X wlx0 : a0 = b0 wlx1 : c0 = d0 wlx2 : e0 = f0 wry0 : b0 = d0 wry1 : a0 = c0 wrz0, wlz0 : d0 = f0 wrz1, wlz1 : c0 = e0 ehrnat_z0 : wrz0 @ 1 = 1 @ wlz0 ehrnat_z1 : wrz1 @ 1 = 1 @ wlz1
forall (wly1 : a0 = c0)
(ehrnat_y1 : wry1 @ 1 = 1 @ wly1),
(fun (q : b0 = d0) (s : wry0 @ 1 = 1 @ q) =>
(fun (q0 : e0 = f0) (s0 : wlx2 @ 1 = 1 @ q0) =>
(fun (q1 : c0 = d0) (s1 : wlx1 @ 1 = 1 @ q1) =>
(fun (q2 : a0 = b0) (s2 : wlx0 @ 1 = 1 @ q2) =>
forall (wlrnat_y_x : wly1 @ q1 = q2 @ q)
(wlrnat_z_x : wlz1 @ q0 = q1 @ wlz0),
whiskerR
((whiskerL wlx0 1 )^ @
((((...)^ @ (...)) @ ((concat_p1 ...)^)^
[I] ((...)^ @ (...)) @ ((concat_p1 ...)^)^) @
whiskerR 1 wlx2)) 1 @
((whiskerR 1 1 )^ @
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 1 ) [-] s0) =
(s2
[-] (whiskerR 1 1 )^ @
((s [-] ehrnat_z0) @ whiskerL 1 1 )) @
whiskerL 1
((whiskerR 1 q0)^ @
((wlrnat_y_x [-] wlrnat_z_x) @ whiskerL q2 1 ))^)
wlx0 (rlucancel 1 )) wlx1 (rlucancel 1 )) wlx2
(rlucancel 1 )) wry0 (rlucancel 1 )
snapply equiv_path_ind_rlucancel. X : Type a0, b0, c0, d0, e0, f0 : X wlx0 : a0 = b0 wlx1 : c0 = d0 wlx2 : e0 = f0 wry0 : b0 = d0 wry1 : a0 = c0 wrz0, wlz0 : d0 = f0 wrz1, wlz1 : c0 = e0 ehrnat_z0 : wrz0 @ 1 = 1 @ wlz0 ehrnat_z1 : wrz1 @ 1 = 1 @ wlz1
(fun (q : a0 = c0) (s : wry1 @ 1 = 1 @ q) =>
(fun (q0 : b0 = d0) (s0 : wry0 @ 1 = 1 @ q0) =>
(fun (q1 : e0 = f0) (s1 : wlx2 @ 1 = 1 @ q1) =>
(fun (q2 : c0 = d0) (s2 : wlx1 @ 1 = 1 @ q2) =>
(fun (q3 : a0 = b0) (s3 : wlx0 @ 1 = 1 @ q3) =>
forall (wlrnat_y_x : q @ q2 = q3 @ q0)
(wlrnat_z_x : wlz1 @ q1 = q2 @ wlz0),
whiskerR
((whiskerL wlx0 1 )^ @
(((...) @ (...^)^ [I] (...) @ (...^)^) @
whiskerR 1 wlx2)) 1 @
((whiskerR 1 1 )^ @
((s [-] ehrnat_z1) @ whiskerL 1 1 ) [-] s1) =
(s3
[-] (whiskerR 1 1 )^ @
((s0 [-] ehrnat_z0) @ whiskerL 1 1 )) @
whiskerL 1
((whiskerR 1 q1)^ @
((wlrnat_y_x [-] wlrnat_z_x) @ whiskerL q3 1 ))^)
wlx0 (rlucancel 1 )) wlx1 (rlucancel 1 )) wlx2
(rlucancel 1 )) wry0 (rlucancel 1 )) wry1
(rlucancel 1 )
revert wlz0 ehrnat_z0.X : Type a0, b0, c0, d0, e0, f0 : X wlx0 : a0 = b0 wlx1 : c0 = d0 wlx2 : e0 = f0 wry0 : b0 = d0 wry1 : a0 = c0 wrz0 : d0 = f0 wrz1, wlz1 : c0 = e0 ehrnat_z1 : wrz1 @ 1 = 1 @ wlz1
forall (wlz0 : d0 = f0)
(ehrnat_z0 : wrz0 @ 1 = 1 @ wlz0),
(fun (q : a0 = c0) (s : wry1 @ 1 = 1 @ q) =>
(fun (q0 : b0 = d0) (s0 : wry0 @ 1 = 1 @ q0) =>
(fun (q1 : e0 = f0) (s1 : wlx2 @ 1 = 1 @ q1) =>
(fun (q2 : c0 = d0) (s2 : wlx1 @ 1 = 1 @ q2) =>
(fun (q3 : a0 = b0) (s3 : wlx0 @ 1 = 1 @ q3) =>
forall (wlrnat_y_x : q @ q2 = q3 @ q0)
(wlrnat_z_x : wlz1 @ q1 = q2 @ wlz0),
whiskerR
((whiskerL wlx0 1 )^ @
((... @ ...^ [I] ... @ ...^) @ whiskerR 1 wlx2))
1 @
((whiskerR 1 1 )^ @
((s [-] ehrnat_z1) @ whiskerL 1 1 ) [-] s1) =
(s3
[-] (whiskerR 1 1 )^ @
((s0 [-] ehrnat_z0) @ whiskerL 1 1 )) @
whiskerL 1
((whiskerR 1 q1)^ @
((wlrnat_y_x [-] wlrnat_z_x) @ whiskerL q3 1 ))^)
wlx0 (rlucancel 1 )) wlx1 (rlucancel 1 )) wlx2
(rlucancel 1 )) wry0 (rlucancel 1 )) wry1
(rlucancel 1 )
snapply equiv_path_ind_rlucancel. X : Type a0, b0, c0, d0, e0, f0 : X wlx0 : a0 = b0 wlx1 : c0 = d0 wlx2 : e0 = f0 wry0 : b0 = d0 wry1 : a0 = c0 wrz0 : d0 = f0 wrz1, wlz1 : c0 = e0 ehrnat_z1 : wrz1 @ 1 = 1 @ wlz1
(fun (q : d0 = f0) (s : wrz0 @ 1 = 1 @ q) =>
(fun (q0 : a0 = c0) (s0 : wry1 @ 1 = 1 @ q0) =>
(fun (q1 : b0 = d0) (s1 : wry0 @ 1 = 1 @ q1) =>
(fun (q2 : e0 = f0) (s2 : wlx2 @ 1 = 1 @ q2) =>
(fun (q3 : c0 = d0) (s3 : wlx1 @ 1 = 1 @ q3) =>
(fun (q4 : a0 = b0) (s4 : wlx0 @ 1 = 1 @ q4) =>
forall (wlrnat_y_x : q0 @ q3 = q4 @ q1)
(wlrnat_z_x : wlz1 @ q2 = q3 @ q),
whiskerR
((whiskerL wlx0 1 )^ @
((...) @ whiskerR 1 wlx2)) 1 @
((whiskerR 1 1 )^ @ ((...) @ whiskerL 1 1 ) [-] s2) =
(s4 [-] (whiskerR 1 1 )^ @ ((...) @ whiskerL 1 1 )) @
whiskerL 1
((whiskerR 1 q2)^ @ ((...) @ whiskerL q4 1 ))^)
wlx0 (rlucancel 1 )) wlx1 (rlucancel 1 )) wlx2
(rlucancel 1 )) wry0 (rlucancel 1 )) wry1
(rlucancel 1 )) wrz0 (rlucancel 1 )
revert wlz1 ehrnat_z1.X : Type a0, b0, c0, d0, e0, f0 : X wlx0 : a0 = b0 wlx1 : c0 = d0 wlx2 : e0 = f0 wry0 : b0 = d0 wry1 : a0 = c0 wrz0 : d0 = f0 wrz1 : c0 = e0
forall (wlz1 : c0 = e0)
(ehrnat_z1 : wrz1 @ 1 = 1 @ wlz1),
(fun (q : d0 = f0) (s : wrz0 @ 1 = 1 @ q) =>
(fun (q0 : a0 = c0) (s0 : wry1 @ 1 = 1 @ q0) =>
(fun (q1 : b0 = d0) (s1 : wry0 @ 1 = 1 @ q1) =>
(fun (q2 : e0 = f0) (s2 : wlx2 @ 1 = 1 @ q2) =>
(fun (q3 : c0 = d0) (s3 : wlx1 @ 1 = 1 @ q3) =>
(fun (q4 : a0 = b0) (s4 : wlx0 @ 1 = 1 @ q4) =>
forall (wlrnat_y_x : q0 @ q3 = q4 @ q1)
(wlrnat_z_x : wlz1 @ q2 = q3 @ q),
whiskerR ((whiskerL wlx0 1 )^ @ (... @ ...)) 1 @
((whiskerR 1 1 )^ @ (... @ ...) [-] s2) =
(s4 [-] (whiskerR 1 1 )^ @ (... @ ...)) @
whiskerL 1 ((whiskerR 1 q2)^ @ (... @ ...))^)
wlx0 (rlucancel 1 )) wlx1 (rlucancel 1 )) wlx2
(rlucancel 1 )) wry0 (rlucancel 1 )) wry1
(rlucancel 1 )) wrz0 (rlucancel 1 )
snapply equiv_path_ind_rlucancel. X : Type a0, b0, c0, d0, e0, f0 : X wlx0 : a0 = b0 wlx1 : c0 = d0 wlx2 : e0 = f0 wry0 : b0 = d0 wry1 : a0 = c0 wrz0 : d0 = f0 wrz1 : c0 = e0
(fun (q : c0 = e0) (s : wrz1 @ 1 = 1 @ q) =>
(fun (q0 : d0 = f0) (s0 : wrz0 @ 1 = 1 @ q0) =>
(fun (q1 : a0 = c0) (s1 : wry1 @ 1 = 1 @ q1) =>
(fun (q2 : b0 = d0) (s2 : wry0 @ 1 = 1 @ q2) =>
(fun (q3 : e0 = f0) (s3 : wlx2 @ 1 = 1 @ q3) =>
(fun (q4 : c0 = d0) (s4 : wlx1 @ 1 = 1 @ q4) =>
(fun (q5 : a0 = b0) (s5 : wlx0 @ 1 = 1 @ q5) =>
forall (wlrnat_y_x : q1 @ q4 = q5 @ q2)
(wlrnat_z_x : q @ q3 = q4 @ q0),
whiskerR (...^ @ ...) 1 @ (...^ @ ... [-] s3) =
(s5 [-] ...^ @ ...) @ whiskerL 1 (...^ @ ...)^)
wlx0 (rlucancel 1 )) wlx1 (rlucancel 1 )) wlx2
(rlucancel 1 )) wry0 (rlucancel 1 )) wry1
(rlucancel 1 )) wrz0 (rlucancel 1 )) wrz1
(rlucancel 1 )
destruct wry0, wry1, wrz0, wrz1.X : Type a0, b0 : X wlx0, wlx1, wlx2 : a0 = b0 wry0 : b0 = b0 wry1 : a0 = a0 wrz0 : b0 = b0 wrz1 : a0 = a0
forall (wlrnat_y_x : 1 @ wlx1 = wlx0 @ 1 )
(wlrnat_z_x : 1 @ wlx2 = wlx1 @ 1 ),
whiskerR
((whiskerL wlx0 1 )^ @
((((concat_p1 (wlx0 @ 1 ))^ @
(((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 wlrnat_y_x^) @
(rlucancel 1 [-] rlucancel 1 )^)) @
((concat_p1 (1 @ wlx1))^)^
[I] ((concat_p1 (wlx1 @ 1 ))^ @
(((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 wlrnat_z_x^) @
(rlucancel 1 [-] rlucancel 1 )^)) @
((concat_p1 (1 @ wlx2))^)^) @ whiskerR 1 wlx2))
1 @
((whiskerR 1 1 )^ @
((rlucancel 1 [-] rlucancel 1 ) @ whiskerL 1 1 )
[-] rlucancel 1 ) =
(rlucancel 1
[-] (whiskerR 1 1 )^ @
((rlucancel 1 [-] rlucancel 1 ) @ whiskerL 1 1 )) @
whiskerL 1
((whiskerR 1 wlx2)^ @
((wlrnat_y_x [-] wlrnat_z_x) @ whiskerL wlx0 1 ))^
clear wry0 wry1 wrz0 wrz1.X : Type a0, b0 : X wlx0, wlx1, wlx2 : a0 = b0
forall (wlrnat_y_x : 1 @ wlx1 = wlx0 @ 1 )
(wlrnat_z_x : 1 @ wlx2 = wlx1 @ 1 ),
whiskerR
((whiskerL wlx0 1 )^ @
((((concat_p1 (wlx0 @ 1 ))^ @
(((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 wlrnat_y_x^) @
(rlucancel 1 [-] rlucancel 1 )^)) @
((concat_p1 (1 @ wlx1))^)^
[I] ((concat_p1 (wlx1 @ 1 ))^ @
(((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 wlrnat_z_x^) @
(rlucancel 1 [-] rlucancel 1 )^)) @
((concat_p1 (1 @ wlx2))^)^) @ whiskerR 1 wlx2))
1 @
((whiskerR 1 1 )^ @
((rlucancel 1 [-] rlucancel 1 ) @ whiskerL 1 1 )
[-] rlucancel 1 ) =
(rlucancel 1
[-] (whiskerR 1 1 )^ @
((rlucancel 1 [-] rlucancel 1 ) @ whiskerL 1 1 )) @
whiskerL 1
((whiskerR 1 wlx2)^ @
((wlrnat_y_x [-] wlrnat_z_x) @ whiskerL wlx0 1 ))^
revert wlx0.X : Type a0, b0 : X wlx1, wlx2 : a0 = b0
forall (wlx0 : a0 = b0)
(wlrnat_y_x : 1 @ wlx1 = wlx0 @ 1 )
(wlrnat_z_x : 1 @ wlx2 = wlx1 @ 1 ),
whiskerR
((whiskerL wlx0 1 )^ @
((((concat_p1 (wlx0 @ 1 ))^ @
(((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 wlrnat_y_x^) @
(rlucancel 1 [-] rlucancel 1 )^)) @
((concat_p1 (1 @ wlx1))^)^
[I] ((concat_p1 (wlx1 @ 1 ))^ @
(((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 wlrnat_z_x^) @
(rlucancel 1 [-] rlucancel 1 )^)) @
((concat_p1 (1 @ wlx2))^)^) @ whiskerR 1 wlx2))
1 @
((whiskerR 1 1 )^ @
((rlucancel 1 [-] rlucancel 1 ) @ whiskerL 1 1 )
[-] rlucancel 1 ) =
(rlucancel 1
[-] (whiskerR 1 1 )^ @
((rlucancel 1 [-] rlucancel 1 ) @ whiskerL 1 1 )) @
whiskerL 1
((whiskerR 1 wlx2)^ @
((wlrnat_y_x [-] wlrnat_z_x) @ whiskerL wlx0 1 ))^
snapply equiv_path_ind_lrucancel. X : Type a0, b0 : X wlx1, wlx2 : a0 = b0
(fun (q : a0 = b0) (s : 1 @ wlx1 = q @ 1 ) =>
forall wlrnat_z_x : 1 @ wlx2 = wlx1 @ 1 ,
whiskerR
((whiskerL q 1 )^ @
((((concat_p1 (q @ 1 ))^ @
(((rlucancel 1 [-] rlucancel 1 ) @ whiskerL 1 s^) @
(rlucancel 1 [-] rlucancel 1 )^)) @
((concat_p1 (1 @ wlx1))^)^
[I] ((concat_p1 (wlx1 @ 1 ))^ @
(((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 wlrnat_z_x^) @
(rlucancel 1 [-] rlucancel 1 )^)) @
((concat_p1 (1 @ wlx2))^)^) @
whiskerR 1 wlx2)) 1 @
((whiskerR 1 1 )^ @
((rlucancel 1 [-] rlucancel 1 ) @ whiskerL 1 1 )
[-] rlucancel 1 ) =
(rlucancel 1
[-] (whiskerR 1 1 )^ @
((rlucancel 1 [-] rlucancel 1 ) @ whiskerL 1 1 )) @
whiskerL 1
((whiskerR 1 wlx2)^ @
((s [-] wlrnat_z_x) @ whiskerL q 1 ))^) wlx1
(lrucancel 1 )
revert wlx1.X : Type a0, b0 : X wlx2 : a0 = b0
forall wlx1 : a0 = b0,
(fun (q : a0 = b0) (s : 1 @ wlx1 = q @ 1 ) =>
forall wlrnat_z_x : 1 @ wlx2 = wlx1 @ 1 ,
whiskerR
((whiskerL q 1 )^ @
((((concat_p1 (q @ 1 ))^ @
(((rlucancel 1 [-] rlucancel 1 ) @ whiskerL 1 s^) @
(rlucancel 1 [-] rlucancel 1 )^)) @
((concat_p1 (1 @ wlx1))^)^
[I] ((concat_p1 (wlx1 @ 1 ))^ @
(((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 wlrnat_z_x^) @
(rlucancel 1 [-] rlucancel 1 )^)) @
((concat_p1 (1 @ wlx2))^)^) @
whiskerR 1 wlx2)) 1 @
((whiskerR 1 1 )^ @
((rlucancel 1 [-] rlucancel 1 ) @ whiskerL 1 1 )
[-] rlucancel 1 ) =
(rlucancel 1
[-] (whiskerR 1 1 )^ @
((rlucancel 1 [-] rlucancel 1 ) @ whiskerL 1 1 )) @
whiskerL 1
((whiskerR 1 wlx2)^ @
((s [-] wlrnat_z_x) @ whiskerL q 1 ))^) wlx1
(lrucancel 1 )
snapply equiv_path_ind_lrucancel. X : Type a0, b0 : X wlx2 : a0 = b0
(fun (q : a0 = b0) (s : 1 @ wlx2 = q @ 1 ) =>
whiskerR
((whiskerL q 1 )^ @
((((concat_p1 (q @ 1 ))^ @
(((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 (lrucancel 1 )^) @
(rlucancel 1 [-] rlucancel 1 )^)) @
((concat_p1 (1 @ q))^)^
[I] ((concat_p1 (q @ 1 ))^ @
(((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 s^) @
(rlucancel 1 [-] rlucancel 1 )^)) @
((concat_p1 (1 @ wlx2))^)^) @
whiskerR 1 wlx2)) 1 @
((whiskerR 1 1 )^ @
((rlucancel 1 [-] rlucancel 1 ) @ whiskerL 1 1 )
[-] rlucancel 1 ) =
(rlucancel 1
[-] (whiskerR 1 1 )^ @
((rlucancel 1 [-] rlucancel 1 ) @ whiskerL 1 1 )) @
whiskerL 1
((whiskerR 1 wlx2)^ @
((lrucancel 1 [-] s) @ whiskerL q 1 ))^) wlx2
(lrucancel 1 )
destruct wlx2.X : Type a0 : X wlx2 : a0 = a0
whiskerR
((whiskerL 1 1 )^ @
((((concat_p1 (1 @ 1 ))^ @
(((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 (lrucancel 1 )^) @
(rlucancel 1 [-] rlucancel 1 )^)) @
((concat_p1 (1 @ 1 ))^)^
[I] ((concat_p1 (1 @ 1 ))^ @
(((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 (lrucancel 1 )^) @
(rlucancel 1 [-] rlucancel 1 )^)) @
((concat_p1 (1 @ 1 ))^)^) @ whiskerR 1 1 )) 1 @
((whiskerR 1 1 )^ @
((rlucancel 1 [-] rlucancel 1 ) @ whiskerL 1 1 )
[-] rlucancel 1 ) =
(rlucancel 1
[-] (whiskerR 1 1 )^ @
((rlucancel 1 [-] rlucancel 1 ) @ whiskerL 1 1 )) @
whiskerL 1
((whiskerR 1 1 )^ @
((lrucancel 1 [-] lrucancel 1 ) @ whiskerL 1 1 ))^
clear wlx2.X : Type a0 : X
whiskerR
((whiskerL 1 1 )^ @
((((concat_p1 (1 @ 1 ))^ @
(((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 (lrucancel 1 )^) @
(rlucancel 1 [-] rlucancel 1 )^)) @
((concat_p1 (1 @ 1 ))^)^
[I] ((concat_p1 (1 @ 1 ))^ @
(((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 (lrucancel 1 )^) @
(rlucancel 1 [-] rlucancel 1 )^)) @
((concat_p1 (1 @ 1 ))^)^) @ whiskerR 1 1 )) 1 @
((whiskerR 1 1 )^ @
((rlucancel 1 [-] rlucancel 1 ) @ whiskerL 1 1 )
[-] rlucancel 1 ) =
(rlucancel 1
[-] (whiskerR 1 1 )^ @
((rlucancel 1 [-] rlucancel 1 ) @ whiskerL 1 1 )) @
whiskerL 1
((whiskerR 1 1 )^ @
((lrucancel 1 [-] lrucancel 1 ) @ whiskerL 1 1 ))^
reflexivity .
Defined .
End wlrnat_V_p_pp .
Definition wlrnat_V_p_pp {X } {a : X} {u v w : idpath a = idpath a} (p : 1 = w) (q : u = 1 ) (r : 1 = v) :
Wlrnat_V_p_pp (ehrnat_pp q r _) (ehrnat_pp q r _) (wlrnat_p_pp p q r) (wlrnat_pp_p q r p)
(wlrnat_V p q) (wlrnat_V p r) =
wlrnat_V p (q @ r).X : Type a : X u, v, w : 1 = 1 p : 1 = wq : u = 1 r : 1 = v
Wlrnat_V_p_pp (ehrnat_pp q r w) (ehrnat_pp q r 1 )
(wlrnat_p_pp p q r) (wlrnat_pp_p q r p)
(wlrnat_V p q) (wlrnat_V p r) = wlrnat_V p (q @ r)
Proof .X : Type a : X u, v, w : 1 = 1 p : 1 = wq : u = 1 r : 1 = v
Wlrnat_V_p_pp (ehrnat_pp q r w) (ehrnat_pp q r 1 )
(wlrnat_p_pp p q r) (wlrnat_pp_p q r p)
(wlrnat_V p q) (wlrnat_V p r) = wlrnat_V p (q @ r)
revert u q.X : Type a : X v, w : 1 = 1 p : 1 = wr : 1 = v
forall (u : 1 = 1 ) (q : u = 1 ),
Wlrnat_V_p_pp (ehrnat_pp q r w) (ehrnat_pp q r 1 )
(wlrnat_p_pp p q r) (wlrnat_pp_p q r p)
(wlrnat_V p q) (wlrnat_V p r) = wlrnat_V p (q @ r)
snapply (equiv_path_ind (equiv_path_inverse _)). X : Type a : X v, w : 1 = 1 p : 1 = wr : 1 = v
(fun (b : 1 = 1 ) (x : (fun y : 1 = 1 => y = 1 ) b) =>
Wlrnat_V_p_pp (ehrnat_pp x r w) (ehrnat_pp x r 1 )
(wlrnat_p_pp p x r) (wlrnat_pp_p x r p)
(wlrnat_V p x) (wlrnat_V p r) = wlrnat_V p (x @ r))
1 (equiv_path_inverse 1 1 1 )
by destruct p, r.
Defined .
(** Next we prove a coherence law relating [eh_V p (q @ r)] to [eh_V p q] and [eh_V p q]. *)
(* The following tactics will be used to make the proof faster, but with only minor modifications, the proof goes through without these tactics. The final tactic [generalize_goal] takes a goal of the form [forall a b c ..., expression] and asserts a new goal [forall P, _ -> forall a b c ..., P a b c ...] which can be used to prove the original goal. Because [expression] has been replaced with a generic function, the proof of the new goal can be more efficient than the proof of the special case, especially when there are around 84 variables. *)
Ltac apply_P ty P :=
lazymatch ty with
| forall a : ?A , ?ty
=> let ty' := fresh in
let P' := fresh in
constr :(forall a : A,
(* Bind [ty] in [match] so that we avoid issues such as https://github.com/coq/coq/issues/7299 and similar ones. Without [return _], [match] tries two ways to elaborate the branches, which results in exponential blowup on failure. *)
match ty, P a return _ with
| ty', P'
=> ltac :(let ty := (eval cbv delta [ty'] in ty') in
let P := (eval cbv delta [P'] in P') in
clear ty' P';
let res := apply_P ty P in
exact res)
end )
| _ => P
end .
Ltac make_P_and_evar ty :=
let P := fresh "P" in
open_constr:(forall P : _, _ -> ltac :(let res := apply_P ty P in exact res)).
Ltac generalize_goal X :=
match goal with |- ?G => let T := make_P_and_evar G in assert (X : T) end .
(* We need this equivalence twice below. *)
Local Lemma equiv_helper {X } {a b : X} {p q r : a = b} (t : q @ 1 = r) (u : p @ 1 = r) (s : p = q)
: ((concat_p1 p)^ @ (u @ t^)) @ (concat_p1 q) = s
<~> whiskerR s 1 @ t = u.X : Type a, b : X p, q, r : a = b t : q @ 1 = r u : p @ 1 = r s : p = q
((concat_p1 p)^ @ (u @ t^)) @ concat_p1 q = s <~>
whiskerR s 1 @ t = u
Proof .X : Type a, b : X p, q, r : a = b t : q @ 1 = r u : p @ 1 = r s : p = q
((concat_p1 p)^ @ (u @ t^)) @ concat_p1 q = s <~>
whiskerR s 1 @ t = u
snapply (_ oE equiv_path_inverse _ _). X : Type a, b : X p, q, r : a = b t : q @ 1 = r u : p @ 1 = r s : p = q
s = ((concat_p1 p)^ @ (u @ t^)) @ concat_p1 q <~>
whiskerR s 1 @ t = u
snapply (_ oE equiv_moveR_pV _ _ _). X : Type a, b : X p, q, r : a = b t : q @ 1 = r u : p @ 1 = r s : p = q
s @ (concat_p1 q)^ = (concat_p1 p)^ @ (u @ t^) <~>
whiskerR s 1 @ t = u
snapply (_ oE equiv_moveR_Mp _ _ _). X : Type a, b : X p, q, r : a = b t : q @ 1 = r u : p @ 1 = r s : p = q
concat_p1 p @ (s @ (concat_p1 q)^) = u @ t^ <~>
whiskerR s 1 @ t = u
snapply (_ oE equiv_concat_l _ _). X : Type a, b : X p, q, r : a = b t : q @ 1 = r u : p @ 1 = r s : p = q
?x = u @ t^ <~> whiskerR s 1 @ t = u
3 : exact (moveL_Mp _ _ _ (moveL_pV _ _ _ (whiskerR_p1 s))).X : Type a, b : X p, q, r : a = b t : q @ 1 = r u : p @ 1 = r s : p = q
whiskerR s 1 = u @ t^ <~> whiskerR s 1 @ t = u
exact (equiv_moveR_pM _ _ _).
Defined .
(* This special case of [equiv_path_ind] comes up a lot. *)
Definition equiv_path_ind_moveL_Mp {X } (a b c : X) (p : a = c) (r : a = b)
(P : forall (q : b = c), p = r @ q -> Type )
(i : P (r^ @ p) (equiv_moveL_Mp _ _ _ 1 ))
: forall (q : b = c) (s : p = r @ q), P q s.X : Type a, b, c : X p : a = c r : a = b P : forall q : b = c, p = r @ q -> Type i : P (r^ @ p) (equiv_moveL_Mp (r^ @ p) p r 1 )
forall (q : b = c) (s : p = r @ q), P q s
Proof .X : Type a, b, c : X p : a = c r : a = b P : forall q : b = c, p = r @ q -> Type i : P (r^ @ p) (equiv_moveL_Mp (r^ @ p) p r 1 )
forall (q : b = c) (s : p = r @ q), P q s
exact (equiv_path_ind (fun q => (equiv_moveL_Mp q _ _)) P i).
Defined .
(* A form of the coherence we can prove by path induction. *)
Definition eh_V_p_pp_gen {X : Type }
(* 0 -paths *)
{a b c d e f : X}
(* 1 -paths *)
{wlx0 x0 wrx0 : a = b}
{wlx1 x1 wrx1 : c = d}
{wlx2 x2 wrx2 : e = f}
{wly0 y0 wry0 : b = d}
{wly1 y1 wry1 : a = c}
{wlz0 z0 wrz0 : d = f}
{wlz1 z1 wrz1 : c = e}
{wlyz0 wryz0 : b = f}
{wlyz1 wryz1 : a = e}
(* 2 -paths *)
{ulnat_x0 : wlx0 @ 1 = 1 @ x0}
{urnat_x0 : wrx0 @ 1 = 1 @ x0}
{ulnat_x1 : wlx1 @ 1 = 1 @ x1}
{urnat_x1 : wrx1 @ 1 = 1 @ x1}
{ulnat_x2 : wlx2 @ 1 = 1 @ x2}
{urnat_x2 : wrx2 @ 1 = 1 @ x2}
{ulnat_y0 : wly0 @ 1 = 1 @ y0}
{urnat_y0 : wry0 @ 1 = 1 @ y0}
{ulnat_y1 : wly1 @ 1 = 1 @ y1}
{urnat_y1 : wry1 @ 1 = 1 @ y1}
{ulnat_z0 : wlz0 @ 1 = 1 @ z0}
{urnat_z0 : wrz0 @ 1 = 1 @ z0}
{ulnat_z1 : wlz1 @ 1 = 1 @ z1}
{urnat_z1 : wrz1 @ 1 = 1 @ z1}
{ulnat_yz0 : wlyz0 @ 1 = 1 @ (y0 @ z0)}
{urnat_yz0 : wryz0 @ 1 = 1 @ (y0 @ z0)}
{ulnat_yz1 : wlyz1 @ 1 = 1 @ (y1 @ z1)}
{urnat_yz1 : wryz1 @ 1 = 1 @ (y1 @ z1)}
{ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0}
{ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1}
{ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2}
{ehrnat_y0 : wry0 @ 1 = 1 @ wly0}
{ehrnat_y1 : wry1 @ 1 = 1 @ wly1}
{ehrnat_z0 : wrz0 @ 1 = 1 @ wlz0}
{ehrnat_z1 : wrz1 @ 1 = 1 @ wlz1}
{ehrnat_yz0 : wryz0 @ 1 = 1 @ wlyz0}
{ehrnat_yz1 : wryz1 @ 1 = 1 @ wlyz1}
{wlrnat_x_y : wlx0 @ wry0 = wry1 @ wlx1}
{wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0}
{wlrnat_x_z : wlx1 @ wrz0 = wrz1 @ wlx2}
{wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0}
{wlrnat_x_yz : wlx0 @ wryz0 = wryz1 @ wlx2}
{wlrnat_yz_x : wlyz1 @ wrx2 = wrx0 @ wlyz0}
{wrpp_yz0 : wry0 @ wrz0 = wryz0}
{wlpp_yz0 : wly0 @ wlz0 = wlyz0}
{wrpp_yz1 : wry1 @ wrz1 = wryz1}
{wlpp_yz1 : wly1 @ wlz1 = wlyz1}
(* 3 -paths *)
{H_ulnat_yz0 : (ulnat_y0 [-] ulnat_z0) = whiskerR wlpp_yz0 _ @ ulnat_yz0}
{H_urnat_yz0 : (urnat_y0 [-] urnat_z0) = whiskerR wrpp_yz0 _ @ urnat_yz0}
{H_ulnat_yz1 : (ulnat_y1 [-] ulnat_z1) = whiskerR wlpp_yz1 _ @ ulnat_yz1}
{H_urnat_yz1 : (urnat_y1 [-] urnat_z1) = whiskerR wrpp_yz1 _ @ urnat_yz1}
{H_ehrnat_yz0 : (ehrnat_y0 [-] ehrnat_z0) @ whiskerL _ wlpp_yz0 =
whiskerR wrpp_yz0 _ @ ehrnat_yz0}
{H_ehrnat_yz1 : (ehrnat_y1 [-] ehrnat_z1) @ whiskerL _ wlpp_yz1 =
whiskerR wrpp_yz1 _ @ ehrnat_yz1}
{H_wlrnat_x_yz : (wlrnat_x_y [I] wlrnat_x_z) @ whiskerR wrpp_yz1 _ =
whiskerL _ wrpp_yz0 @ wlrnat_x_yz}
{H_wlrnat_yz_x : (wlrnat_y_x [-] wlrnat_z_x) @ whiskerL _ wlpp_yz0 =
whiskerR wlpp_yz1 _ @ wlrnat_yz_x}
(ehlnat_1p_x0 : (ehlnat_x0 [I] urnat_x0) @ 1 = 1 @ ulnat_x0)
(ehlnat_1p_x1 : (ehlnat_x1 [I] urnat_x1) @ 1 = 1 @ ulnat_x1)
(ehlnat_1p_x2 : (ehlnat_x2 [I] urnat_x2) @ 1 = 1 @ ulnat_x2)
{ehrnat_p1_y0 : (ehrnat_y0 [I] ulnat_y0) @ 1 = 1 @ urnat_y0}
{ehrnat_p1_y1 : (ehrnat_y1 [I] ulnat_y1) @ 1 = 1 @ urnat_y1}
{ehrnat_p1_z0 : (ehrnat_z0 [I] ulnat_z0) @ 1 = 1 @ urnat_z0}
{ehrnat_p1_z1 : (ehrnat_z1 [I] ulnat_z1) @ 1 = 1 @ urnat_z1}
{ehrnat_p1_yz0 : (ehrnat_yz0 [I] ulnat_yz0) @ 1 = 1 @ urnat_yz0}
{ehrnat_p1_yz1 : (ehrnat_yz1 [I] ulnat_yz1) @ 1 = 1 @ urnat_yz1}
{wlrnat_V_x_y : whiskerR wlrnat_x_y _ @ (ehrnat_y1 [-] ehlnat_x1) =
(ehlnat_x0 [-] ehrnat_y0) @ whiskerL _ wlrnat_y_x^}
{wlrnat_V_x_z : whiskerR wlrnat_x_z _ @ (ehrnat_z1 [-] ehlnat_x2) =
(ehlnat_x1 [-] ehrnat_z0) @ whiskerL _ wlrnat_z_x^}
{wlrnat_V_x_yz : whiskerR wlrnat_x_yz _ @ (ehrnat_yz1 [-] ehlnat_x2) =
(ehlnat_x0 [-] ehrnat_yz0) @ whiskerL _ wlrnat_yz_x^}
(* 4 -paths *)
(H_ehrnat_p1_yz0 :
Ehrnat_p1_pp 1 1 H_ehrnat_yz0 H_ulnat_yz0 H_urnat_yz0 ehrnat_p1_y0 ehrnat_p1_z0 =
ehrnat_p1_yz0)
(H_ehrnat_p1_yz1 :
Ehrnat_p1_pp 1 1 H_ehrnat_yz1 H_ulnat_yz1 H_urnat_yz1 ehrnat_p1_y1 ehrnat_p1_z1 =
ehrnat_p1_yz1)
(H_wlrnat_V_x_yz :
Wlrnat_V_p_pp H_ehrnat_yz0 H_ehrnat_yz1 H_wlrnat_x_yz H_wlrnat_yz_x wlrnat_V_x_y wlrnat_V_x_z =
wlrnat_V_x_yz)
: let eh_x_y := concat_p_pp x0 y0 z0 @
whiskerR (((rlucancel_inv (ulnat_x0 [-] urnat_y0))^ @ wlrnat_x_y) @
rlucancel_inv (urnat_y1 [-] ulnat_x1)) z0 in
whiskerR (concat_p1 _ @@ concat_p1 _) eh_x_y @
whiskerR (eh_V_gen (ehlnat_1p_x0) (ehlnat_1p_x2) (ehrnat_p1_yz0) (ehrnat_p1_yz1) wlrnat_V_x_yz) eh_x_y @
lrucancel 1 @
whiskerL eh_x_y (Syllepsis.concat_pp_p_p_pp _ _ _)^ @
whiskerL eh_x_y (concat_p1 _ @@ concat_p1 _)^ =
(eh_p_pp_gen H_urnat_yz0 H_urnat_yz1 H_wlrnat_x_yz [-]
lrucancel (whiskerL _ (ap (fun p => whiskerL y1 p)
(moveL_V1 _ _ (eh_V_gen ehlnat_1p_x1 ehlnat_1p_x2
ehrnat_p1_z0 ehrnat_p1_z1 wlrnat_V_x_z))))) [-]
(eh_pp_p_gen H_ulnat_yz1 H_ulnat_yz0 H_wlrnat_yz_x [-]
lrucancel (whiskerL _ (ap (fun p => whiskerR p z0)
(moveL_1V _ _ (eh_V_gen ehlnat_1p_x0 ehlnat_1p_x1
ehrnat_p1_y0 ehrnat_p1_y1 wlrnat_V_x_y))))).X : Type a, b, c, d, e, f : X wlx0, x0, wrx0 : a = b wlx1, x1, wrx1 : c = d wlx2, x2, wrx2 : e = f wly0, y0, wry0 : b = d wly1, y1, wry1 : a = c wlz0, z0, wrz0 : d = f wlz1, z1, wrz1 : c = e wlyz0, wryz0 : b = f wlyz1, wryz1 : a = e ulnat_x0 : wlx0 @ 1 = 1 @ x0 urnat_x0 : wrx0 @ 1 = 1 @ x0 ulnat_x1 : wlx1 @ 1 = 1 @ x1 urnat_x1 : wrx1 @ 1 = 1 @ x1 ulnat_x2 : wlx2 @ 1 = 1 @ x2 urnat_x2 : wrx2 @ 1 = 1 @ x2 ulnat_y0 : wly0 @ 1 = 1 @ y0 urnat_y0 : wry0 @ 1 = 1 @ y0 ulnat_y1 : wly1 @ 1 = 1 @ y1 urnat_y1 : wry1 @ 1 = 1 @ y1 ulnat_z0 : wlz0 @ 1 = 1 @ z0 urnat_z0 : wrz0 @ 1 = 1 @ z0 ulnat_z1 : wlz1 @ 1 = 1 @ z1 urnat_z1 : wrz1 @ 1 = 1 @ z1 ulnat_yz0 : wlyz0 @ 1 = 1 @ (y0 @ z0) urnat_yz0 : wryz0 @ 1 = 1 @ (y0 @ z0) ulnat_yz1 : wlyz1 @ 1 = 1 @ (y1 @ z1) urnat_yz1 : wryz1 @ 1 = 1 @ (y1 @ z1) ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : wry0 @ 1 = 1 @ wly0 ehrnat_y1 : wry1 @ 1 = 1 @ wly1 ehrnat_z0 : wrz0 @ 1 = 1 @ wlz0 ehrnat_z1 : wrz1 @ 1 = 1 @ wlz1 ehrnat_yz0 : wryz0 @ 1 = 1 @ wlyz0 ehrnat_yz1 : wryz1 @ 1 = 1 @ wlyz1 wlrnat_x_y : wlx0 @ wry0 = wry1 @ wlx1 wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_x_z : wlx1 @ wrz0 = wrz1 @ wlx2 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0 wlrnat_x_yz : wlx0 @ wryz0 = wryz1 @ wlx2 wlrnat_yz_x : wlyz1 @ wrx2 = wrx0 @ wlyz0 wrpp_yz0 : wry0 @ wrz0 = wryz0 wlpp_yz0 : wly0 @ wlz0 = wlyz0 wrpp_yz1 : wry1 @ wrz1 = wryz1 wlpp_yz1 : wly1 @ wlz1 = wlyz1 H_ulnat_yz0 : ulnat_y0 [-] ulnat_z0 =
whiskerR wlpp_yz0 1 @ ulnat_yz0 H_urnat_yz0 : urnat_y0 [-] urnat_z0 =
whiskerR wrpp_yz0 1 @ urnat_yz0 H_ulnat_yz1 : ulnat_y1 [-] ulnat_z1 =
whiskerR wlpp_yz1 1 @ ulnat_yz1 H_urnat_yz1 : urnat_y1 [-] urnat_z1 =
whiskerR wrpp_yz1 1 @ urnat_yz1 H_ehrnat_yz0 : (ehrnat_y0 [-] ehrnat_z0) @
whiskerL 1 wlpp_yz0 =
whiskerR wrpp_yz0 1 @ ehrnat_yz0 H_ehrnat_yz1 : (ehrnat_y1 [-] ehrnat_z1) @
whiskerL 1 wlpp_yz1 =
whiskerR wrpp_yz1 1 @ ehrnat_yz1 H_wlrnat_x_yz : (wlrnat_x_y [I] wlrnat_x_z) @
whiskerR wrpp_yz1 wlx2 =
whiskerL wlx0 wrpp_yz0 @ wlrnat_x_yz H_wlrnat_yz_x : (wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 wlpp_yz0 =
whiskerR wlpp_yz1 wrx2 @ wlrnat_yz_x ehlnat_1p_x0 : (ehlnat_x0 [I] urnat_x0) @ 1 =
1 @ ulnat_x0 ehlnat_1p_x1 : (ehlnat_x1 [I] urnat_x1) @ 1 =
1 @ ulnat_x1 ehlnat_1p_x2 : (ehlnat_x2 [I] urnat_x2) @ 1 =
1 @ ulnat_x2 ehrnat_p1_y0 : (ehrnat_y0 [I] ulnat_y0) @ 1 =
1 @ urnat_y0 ehrnat_p1_y1 : (ehrnat_y1 [I] ulnat_y1) @ 1 =
1 @ urnat_y1 ehrnat_p1_z0 : (ehrnat_z0 [I] ulnat_z0) @ 1 =
1 @ urnat_z0 ehrnat_p1_z1 : (ehrnat_z1 [I] ulnat_z1) @ 1 =
1 @ urnat_z1 ehrnat_p1_yz0 : (ehrnat_yz0 [I] ulnat_yz0) @ 1 =
1 @ urnat_yz0 ehrnat_p1_yz1 : (ehrnat_yz1 [I] ulnat_yz1) @ 1 =
1 @ urnat_yz1 wlrnat_V_x_y : whiskerR wlrnat_x_y 1 @
(ehrnat_y1 [-] ehlnat_x1) =
(ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^ wlrnat_V_x_z : whiskerR wlrnat_x_z 1 @
(ehrnat_z1 [-] ehlnat_x2) =
(ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^ wlrnat_V_x_yz : whiskerR wlrnat_x_yz 1 @
(ehrnat_yz1 [-] ehlnat_x2) =
(ehlnat_x0 [-] ehrnat_yz0) @
whiskerL 1 wlrnat_yz_x^ H_ehrnat_p1_yz0 : Ehrnat_p1_pp 1 1 H_ehrnat_yz0
H_ulnat_yz0 H_urnat_yz0
ehrnat_p1_y0 ehrnat_p1_z0 =
ehrnat_p1_yz0 H_ehrnat_p1_yz1 : Ehrnat_p1_pp 1 1 H_ehrnat_yz1
H_ulnat_yz1 H_urnat_yz1
ehrnat_p1_y1 ehrnat_p1_z1 =
ehrnat_p1_yz1 H_wlrnat_V_x_yz : Wlrnat_V_p_pp H_ehrnat_yz0
H_ehrnat_yz1 H_wlrnat_x_yz
H_wlrnat_yz_x wlrnat_V_x_y
wlrnat_V_x_z = wlrnat_V_x_yz
let eh_x_y :=
concat_p_pp x0 y0 z0 @
whiskerR
(((rlucancel_inv (ulnat_x0 [-] urnat_y0))^ @
wlrnat_x_y) @
rlucancel_inv (urnat_y1 [-] ulnat_x1)) z0 in
(((whiskerR
(concat_p1
(((rlucancel_inv (ulnat_x0 [-] urnat_yz0))^ @
wlrnat_x_yz) @
rlucancel_inv (urnat_yz1 [-] ulnat_x2)) @@
concat_p1
(((rlucancel_inv (ulnat_yz1 [-] urnat_x2))^ @
wlrnat_yz_x) @
rlucancel_inv (urnat_x0 [-] ulnat_yz0)))
eh_x_y @
whiskerR
(eh_V_gen ehlnat_1p_x0 ehlnat_1p_x2 ehrnat_p1_yz0
ehrnat_p1_yz1 wlrnat_V_x_yz) eh_x_y) @
lrucancel 1 ) @
whiskerL eh_x_y (concat_pp_p_p_pp y1 x1 z0)^) @
whiskerL eh_x_y
(concat_p1 (concat_pp_p y1 x1 z0) @@
concat_p1 (concat_p_pp y1 x1 z0))^ =
(eh_p_pp_gen H_urnat_yz0 H_urnat_yz1 H_wlrnat_x_yz
[-] lrucancel
(whiskerL (concat_pp_p y1 z1 x2)
(ap
(fun p : z1 @ x2 = x1 @ z0 =>
whiskerL y1 p)
(moveL_V1
(((rlucancel_inv
(ulnat_z1 [-] urnat_x2))^ @
wlrnat_z_x) @
rlucancel_inv (urnat_x1 [-] ulnat_z0))
(((rlucancel_inv
(ulnat_x1 [-] urnat_z0))^ @
wlrnat_x_z) @
rlucancel_inv (urnat_z1 [-] ulnat_x2))
(eh_V_gen ehlnat_1p_x1 ehlnat_1p_x2
ehrnat_p1_z0 ehrnat_p1_z1
wlrnat_V_x_z)))))
[-] (eh_pp_p_gen H_ulnat_yz1 H_ulnat_yz0 H_wlrnat_yz_x
[-] lrucancel
(whiskerL (concat_p_pp x0 y0 z0)
(ap
(fun p : x0 @ y0 = y1 @ x1 =>
whiskerR p z0)
(moveL_1V
(((rlucancel_inv
(ulnat_x0 [-] urnat_y0))^ @
wlrnat_x_y) @
rlucancel_inv
(urnat_y1 [-] ulnat_x1))
(((rlucancel_inv
(ulnat_y1 [-] urnat_x1))^ @
wlrnat_y_x) @
rlucancel_inv
(urnat_x0 [-] ulnat_y0))
(eh_V_gen ehlnat_1p_x0
ehlnat_1p_x1 ehrnat_p1_y0
ehrnat_p1_y1 wlrnat_V_x_y)))))
Proof .X : Type a, b, c, d, e, f : X wlx0, x0, wrx0 : a = b wlx1, x1, wrx1 : c = d wlx2, x2, wrx2 : e = f wly0, y0, wry0 : b = d wly1, y1, wry1 : a = c wlz0, z0, wrz0 : d = f wlz1, z1, wrz1 : c = e wlyz0, wryz0 : b = f wlyz1, wryz1 : a = e ulnat_x0 : wlx0 @ 1 = 1 @ x0 urnat_x0 : wrx0 @ 1 = 1 @ x0 ulnat_x1 : wlx1 @ 1 = 1 @ x1 urnat_x1 : wrx1 @ 1 = 1 @ x1 ulnat_x2 : wlx2 @ 1 = 1 @ x2 urnat_x2 : wrx2 @ 1 = 1 @ x2 ulnat_y0 : wly0 @ 1 = 1 @ y0 urnat_y0 : wry0 @ 1 = 1 @ y0 ulnat_y1 : wly1 @ 1 = 1 @ y1 urnat_y1 : wry1 @ 1 = 1 @ y1 ulnat_z0 : wlz0 @ 1 = 1 @ z0 urnat_z0 : wrz0 @ 1 = 1 @ z0 ulnat_z1 : wlz1 @ 1 = 1 @ z1 urnat_z1 : wrz1 @ 1 = 1 @ z1 ulnat_yz0 : wlyz0 @ 1 = 1 @ (y0 @ z0) urnat_yz0 : wryz0 @ 1 = 1 @ (y0 @ z0) ulnat_yz1 : wlyz1 @ 1 = 1 @ (y1 @ z1) urnat_yz1 : wryz1 @ 1 = 1 @ (y1 @ z1) ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : wry0 @ 1 = 1 @ wly0 ehrnat_y1 : wry1 @ 1 = 1 @ wly1 ehrnat_z0 : wrz0 @ 1 = 1 @ wlz0 ehrnat_z1 : wrz1 @ 1 = 1 @ wlz1 ehrnat_yz0 : wryz0 @ 1 = 1 @ wlyz0 ehrnat_yz1 : wryz1 @ 1 = 1 @ wlyz1 wlrnat_x_y : wlx0 @ wry0 = wry1 @ wlx1 wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_x_z : wlx1 @ wrz0 = wrz1 @ wlx2 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0 wlrnat_x_yz : wlx0 @ wryz0 = wryz1 @ wlx2 wlrnat_yz_x : wlyz1 @ wrx2 = wrx0 @ wlyz0 wrpp_yz0 : wry0 @ wrz0 = wryz0 wlpp_yz0 : wly0 @ wlz0 = wlyz0 wrpp_yz1 : wry1 @ wrz1 = wryz1 wlpp_yz1 : wly1 @ wlz1 = wlyz1 H_ulnat_yz0 : ulnat_y0 [-] ulnat_z0 =
whiskerR wlpp_yz0 1 @ ulnat_yz0 H_urnat_yz0 : urnat_y0 [-] urnat_z0 =
whiskerR wrpp_yz0 1 @ urnat_yz0 H_ulnat_yz1 : ulnat_y1 [-] ulnat_z1 =
whiskerR wlpp_yz1 1 @ ulnat_yz1 H_urnat_yz1 : urnat_y1 [-] urnat_z1 =
whiskerR wrpp_yz1 1 @ urnat_yz1 H_ehrnat_yz0 : (ehrnat_y0 [-] ehrnat_z0) @
whiskerL 1 wlpp_yz0 =
whiskerR wrpp_yz0 1 @ ehrnat_yz0 H_ehrnat_yz1 : (ehrnat_y1 [-] ehrnat_z1) @
whiskerL 1 wlpp_yz1 =
whiskerR wrpp_yz1 1 @ ehrnat_yz1 H_wlrnat_x_yz : (wlrnat_x_y [I] wlrnat_x_z) @
whiskerR wrpp_yz1 wlx2 =
whiskerL wlx0 wrpp_yz0 @ wlrnat_x_yz H_wlrnat_yz_x : (wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 wlpp_yz0 =
whiskerR wlpp_yz1 wrx2 @ wlrnat_yz_x ehlnat_1p_x0 : (ehlnat_x0 [I] urnat_x0) @ 1 =
1 @ ulnat_x0 ehlnat_1p_x1 : (ehlnat_x1 [I] urnat_x1) @ 1 =
1 @ ulnat_x1 ehlnat_1p_x2 : (ehlnat_x2 [I] urnat_x2) @ 1 =
1 @ ulnat_x2 ehrnat_p1_y0 : (ehrnat_y0 [I] ulnat_y0) @ 1 =
1 @ urnat_y0 ehrnat_p1_y1 : (ehrnat_y1 [I] ulnat_y1) @ 1 =
1 @ urnat_y1 ehrnat_p1_z0 : (ehrnat_z0 [I] ulnat_z0) @ 1 =
1 @ urnat_z0 ehrnat_p1_z1 : (ehrnat_z1 [I] ulnat_z1) @ 1 =
1 @ urnat_z1 ehrnat_p1_yz0 : (ehrnat_yz0 [I] ulnat_yz0) @ 1 =
1 @ urnat_yz0 ehrnat_p1_yz1 : (ehrnat_yz1 [I] ulnat_yz1) @ 1 =
1 @ urnat_yz1 wlrnat_V_x_y : whiskerR wlrnat_x_y 1 @
(ehrnat_y1 [-] ehlnat_x1) =
(ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^ wlrnat_V_x_z : whiskerR wlrnat_x_z 1 @
(ehrnat_z1 [-] ehlnat_x2) =
(ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^ wlrnat_V_x_yz : whiskerR wlrnat_x_yz 1 @
(ehrnat_yz1 [-] ehlnat_x2) =
(ehlnat_x0 [-] ehrnat_yz0) @
whiskerL 1 wlrnat_yz_x^ H_ehrnat_p1_yz0 : Ehrnat_p1_pp 1 1 H_ehrnat_yz0
H_ulnat_yz0 H_urnat_yz0
ehrnat_p1_y0 ehrnat_p1_z0 =
ehrnat_p1_yz0 H_ehrnat_p1_yz1 : Ehrnat_p1_pp 1 1 H_ehrnat_yz1
H_ulnat_yz1 H_urnat_yz1
ehrnat_p1_y1 ehrnat_p1_z1 =
ehrnat_p1_yz1 H_wlrnat_V_x_yz : Wlrnat_V_p_pp H_ehrnat_yz0
H_ehrnat_yz1 H_wlrnat_x_yz
H_wlrnat_yz_x wlrnat_V_x_y
wlrnat_V_x_z = wlrnat_V_x_yz
let eh_x_y :=
concat_p_pp x0 y0 z0 @
whiskerR
(((rlucancel_inv (ulnat_x0 [-] urnat_y0))^ @
wlrnat_x_y) @
rlucancel_inv (urnat_y1 [-] ulnat_x1)) z0 in
(((whiskerR
(concat_p1
(((rlucancel_inv (ulnat_x0 [-] urnat_yz0))^ @
wlrnat_x_yz) @
rlucancel_inv (urnat_yz1 [-] ulnat_x2)) @@
concat_p1
(((rlucancel_inv (ulnat_yz1 [-] urnat_x2))^ @
wlrnat_yz_x) @
rlucancel_inv (urnat_x0 [-] ulnat_yz0)))
eh_x_y @
whiskerR
(eh_V_gen ehlnat_1p_x0 ehlnat_1p_x2 ehrnat_p1_yz0
ehrnat_p1_yz1 wlrnat_V_x_yz) eh_x_y) @
lrucancel 1 ) @
whiskerL eh_x_y (concat_pp_p_p_pp y1 x1 z0)^) @
whiskerL eh_x_y
(concat_p1 (concat_pp_p y1 x1 z0) @@
concat_p1 (concat_p_pp y1 x1 z0))^ =
(eh_p_pp_gen H_urnat_yz0 H_urnat_yz1 H_wlrnat_x_yz
[-] lrucancel
(whiskerL (concat_pp_p y1 z1 x2)
(ap
(fun p : z1 @ x2 = x1 @ z0 =>
whiskerL y1 p)
(moveL_V1
(((rlucancel_inv
(ulnat_z1 [-] urnat_x2))^ @
wlrnat_z_x) @
rlucancel_inv (urnat_x1 [-] ulnat_z0))
(((rlucancel_inv
(ulnat_x1 [-] urnat_z0))^ @
wlrnat_x_z) @
rlucancel_inv (urnat_z1 [-] ulnat_x2))
(eh_V_gen ehlnat_1p_x1 ehlnat_1p_x2
ehrnat_p1_z0 ehrnat_p1_z1
wlrnat_V_x_z)))))
[-] (eh_pp_p_gen H_ulnat_yz1 H_ulnat_yz0 H_wlrnat_yz_x
[-] lrucancel
(whiskerL (concat_p_pp x0 y0 z0)
(ap
(fun p : x0 @ y0 = y1 @ x1 =>
whiskerR p z0)
(moveL_1V
(((rlucancel_inv
(ulnat_x0 [-] urnat_y0))^ @
wlrnat_x_y) @
rlucancel_inv
(urnat_y1 [-] ulnat_x1))
(((rlucancel_inv
(ulnat_y1 [-] urnat_x1))^ @
wlrnat_y_x) @
rlucancel_inv
(urnat_x0 [-] ulnat_y0))
(eh_V_gen ehlnat_1p_x0
ehlnat_1p_x1 ehrnat_p1_y0
ehrnat_p1_y1 wlrnat_V_x_y)))))
(* For some reason, it's most efficient to destruct a few things here but the rest within the subgoal. *)
destruct H_ehrnat_p1_yz0, H_ehrnat_p1_yz1, H_wlrnat_V_x_yz.X : Type a, b, c, d, e, f : X wlx0, x0, wrx0 : a = b wlx1, x1, wrx1 : c = d wlx2, x2, wrx2 : e = f wly0, y0, wry0 : b = d wly1, y1, wry1 : a = c wlz0, z0, wrz0 : d = f wlz1, z1, wrz1 : c = e wlyz0, wryz0 : b = f wlyz1, wryz1 : a = e ulnat_x0 : wlx0 @ 1 = 1 @ x0 urnat_x0 : wrx0 @ 1 = 1 @ x0 ulnat_x1 : wlx1 @ 1 = 1 @ x1 urnat_x1 : wrx1 @ 1 = 1 @ x1 ulnat_x2 : wlx2 @ 1 = 1 @ x2 urnat_x2 : wrx2 @ 1 = 1 @ x2 ulnat_y0 : wly0 @ 1 = 1 @ y0 urnat_y0 : wry0 @ 1 = 1 @ y0 ulnat_y1 : wly1 @ 1 = 1 @ y1 urnat_y1 : wry1 @ 1 = 1 @ y1 ulnat_z0 : wlz0 @ 1 = 1 @ z0 urnat_z0 : wrz0 @ 1 = 1 @ z0 ulnat_z1 : wlz1 @ 1 = 1 @ z1 urnat_z1 : wrz1 @ 1 = 1 @ z1 ulnat_yz0 : wlyz0 @ 1 = 1 @ (y0 @ z0) urnat_yz0 : wryz0 @ 1 = 1 @ (y0 @ z0) ulnat_yz1 : wlyz1 @ 1 = 1 @ (y1 @ z1) urnat_yz1 : wryz1 @ 1 = 1 @ (y1 @ z1) ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : wry0 @ 1 = 1 @ wly0 ehrnat_y1 : wry1 @ 1 = 1 @ wly1 ehrnat_z0 : wrz0 @ 1 = 1 @ wlz0 ehrnat_z1 : wrz1 @ 1 = 1 @ wlz1 ehrnat_yz0 : wryz0 @ 1 = 1 @ wlyz0 ehrnat_yz1 : wryz1 @ 1 = 1 @ wlyz1 wlrnat_x_y : wlx0 @ wry0 = wry1 @ wlx1 wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_x_z : wlx1 @ wrz0 = wrz1 @ wlx2 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0 wlrnat_x_yz : wlx0 @ wryz0 = wryz1 @ wlx2 wlrnat_yz_x : wlyz1 @ wrx2 = wrx0 @ wlyz0 wrpp_yz0 : wry0 @ wrz0 = wryz0 wlpp_yz0 : wly0 @ wlz0 = wlyz0 wrpp_yz1 : wry1 @ wrz1 = wryz1 wlpp_yz1 : wly1 @ wlz1 = wlyz1 H_ulnat_yz0 : ulnat_y0 [-] ulnat_z0 =
whiskerR wlpp_yz0 1 @ ulnat_yz0 H_urnat_yz0 : urnat_y0 [-] urnat_z0 =
whiskerR wrpp_yz0 1 @ urnat_yz0 H_ulnat_yz1 : ulnat_y1 [-] ulnat_z1 =
whiskerR wlpp_yz1 1 @ ulnat_yz1 H_urnat_yz1 : urnat_y1 [-] urnat_z1 =
whiskerR wrpp_yz1 1 @ urnat_yz1 H_ehrnat_yz0 : (ehrnat_y0 [-] ehrnat_z0) @
whiskerL 1 wlpp_yz0 =
whiskerR wrpp_yz0 1 @ ehrnat_yz0 H_ehrnat_yz1 : (ehrnat_y1 [-] ehrnat_z1) @
whiskerL 1 wlpp_yz1 =
whiskerR wrpp_yz1 1 @ ehrnat_yz1 H_wlrnat_x_yz : (wlrnat_x_y [I] wlrnat_x_z) @
whiskerR wrpp_yz1 wlx2 =
whiskerL wlx0 wrpp_yz0 @ wlrnat_x_yz H_wlrnat_yz_x : (wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 wlpp_yz0 =
whiskerR wlpp_yz1 wrx2 @ wlrnat_yz_x ehlnat_1p_x0 : (ehlnat_x0 [I] urnat_x0) @ 1 =
1 @ ulnat_x0 ehlnat_1p_x1 : (ehlnat_x1 [I] urnat_x1) @ 1 =
1 @ ulnat_x1 ehlnat_1p_x2 : (ehlnat_x2 [I] urnat_x2) @ 1 =
1 @ ulnat_x2 ehrnat_p1_y0 : (ehrnat_y0 [I] ulnat_y0) @ 1 =
1 @ urnat_y0 ehrnat_p1_y1 : (ehrnat_y1 [I] ulnat_y1) @ 1 =
1 @ urnat_y1 ehrnat_p1_z0 : (ehrnat_z0 [I] ulnat_z0) @ 1 =
1 @ urnat_z0 ehrnat_p1_z1 : (ehrnat_z1 [I] ulnat_z1) @ 1 =
1 @ urnat_z1 wlrnat_V_x_y : whiskerR wlrnat_x_y 1 @
(ehrnat_y1 [-] ehlnat_x1) =
(ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^ wlrnat_V_x_z : whiskerR wlrnat_x_z 1 @
(ehrnat_z1 [-] ehlnat_x2) =
(ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^
let eh_x_y :=
concat_p_pp x0 y0 z0 @
whiskerR
(((rlucancel_inv (ulnat_x0 [-] urnat_y0))^ @
wlrnat_x_y) @
rlucancel_inv (urnat_y1 [-] ulnat_x1)) z0 in
(((whiskerR
(concat_p1
(((rlucancel_inv (ulnat_x0 [-] urnat_yz0))^ @
wlrnat_x_yz) @
rlucancel_inv (urnat_yz1 [-] ulnat_x2)) @@
concat_p1
(((rlucancel_inv (ulnat_yz1 [-] urnat_x2))^ @
wlrnat_yz_x) @
rlucancel_inv (urnat_x0 [-] ulnat_yz0)))
eh_x_y @
whiskerR
(eh_V_gen ehlnat_1p_x0 ehlnat_1p_x2
(Ehrnat_p1_pp 1 1 H_ehrnat_yz0 H_ulnat_yz0
H_urnat_yz0 ehrnat_p1_y0 ehrnat_p1_z0)
(Ehrnat_p1_pp 1 1 H_ehrnat_yz1 H_ulnat_yz1
H_urnat_yz1 ehrnat_p1_y1 ehrnat_p1_z1)
(Wlrnat_V_p_pp H_ehrnat_yz0 H_ehrnat_yz1
H_wlrnat_x_yz H_wlrnat_yz_x wlrnat_V_x_y
wlrnat_V_x_z)) eh_x_y) @
lrucancel 1 ) @
whiskerL eh_x_y (concat_pp_p_p_pp y1 x1 z0)^) @
whiskerL eh_x_y
(concat_p1 (concat_pp_p y1 x1 z0) @@
concat_p1 (concat_p_pp y1 x1 z0))^ =
(eh_p_pp_gen H_urnat_yz0 H_urnat_yz1 H_wlrnat_x_yz
[-] lrucancel
(whiskerL (concat_pp_p y1 z1 x2)
(ap
(fun p : z1 @ x2 = x1 @ z0 =>
whiskerL y1 p)
(moveL_V1
(((rlucancel_inv
(ulnat_z1 [-] urnat_x2))^ @
wlrnat_z_x) @
rlucancel_inv (urnat_x1 [-] ulnat_z0))
(((rlucancel_inv
(ulnat_x1 [-] urnat_z0))^ @
wlrnat_x_z) @
rlucancel_inv (urnat_z1 [-] ulnat_x2))
(eh_V_gen ehlnat_1p_x1 ehlnat_1p_x2
ehrnat_p1_z0 ehrnat_p1_z1
wlrnat_V_x_z)))))
[-] (eh_pp_p_gen H_ulnat_yz1 H_ulnat_yz0 H_wlrnat_yz_x
[-] lrucancel
(whiskerL (concat_p_pp x0 y0 z0)
(ap
(fun p : x0 @ y0 = y1 @ x1 =>
whiskerR p z0)
(moveL_1V
(((rlucancel_inv
(ulnat_x0 [-] urnat_y0))^ @
wlrnat_x_y) @
rlucancel_inv
(urnat_y1 [-] ulnat_x1))
(((rlucancel_inv
(ulnat_y1 [-] urnat_x1))^ @
wlrnat_y_x) @
rlucancel_inv
(urnat_x0 [-] ulnat_y0))
(eh_V_gen ehlnat_1p_x0
ehlnat_1p_x1 ehrnat_p1_y0
ehrnat_p1_y1 wlrnat_V_x_y)))))
(* For efficiency purposes, we generalize the goal to an arbitrary function [P] of the context (except for [X] and [a]), and do all of the induction steps in this generality. This reduces the size of the term that Coq needs to manipulate, speeding up the proof. The same proof works with the next three lines removed and with the second and third last lines removed. *)
revert_until a. X : Type a : X
forall (b c d e f : X) (wlx0 x0 wrx0 : a = b)
(wlx1 x1 wrx1 : c = d) (wlx2 x2 wrx2 : e = f)
(wly0 y0 wry0 : b = d) (wly1 y1 wry1 : a = c)
(wlz0 z0 wrz0 : d = f) (wlz1 z1 wrz1 : c = e)
(wlyz0 wryz0 : b = f) (wlyz1 wryz1 : a = e)
(ulnat_x0 : wlx0 @ 1 = 1 @ x0)
(urnat_x0 : wrx0 @ 1 = 1 @ x0)
(ulnat_x1 : wlx1 @ 1 = 1 @ x1)
(urnat_x1 : wrx1 @ 1 = 1 @ x1)
(ulnat_x2 : wlx2 @ 1 = 1 @ x2)
(urnat_x2 : wrx2 @ 1 = 1 @ x2)
(ulnat_y0 : wly0 @ 1 = 1 @ y0)
(urnat_y0 : wry0 @ 1 = 1 @ y0)
(ulnat_y1 : wly1 @ 1 = 1 @ y1)
(urnat_y1 : wry1 @ 1 = 1 @ y1)
(ulnat_z0 : wlz0 @ 1 = 1 @ z0)
(urnat_z0 : wrz0 @ 1 = 1 @ z0)
(ulnat_z1 : wlz1 @ 1 = 1 @ z1)
(urnat_z1 : wrz1 @ 1 = 1 @ z1)
(ulnat_yz0 : wlyz0 @ 1 = 1 @ (y0 @ z0))
(urnat_yz0 : wryz0 @ 1 = 1 @ (y0 @ z0))
(ulnat_yz1 : wlyz1 @ 1 = 1 @ (y1 @ z1))
(urnat_yz1 : wryz1 @ 1 = 1 @ (y1 @ z1))
(ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0)
(ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1)
(ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2)
(ehrnat_y0 : wry0 @ 1 = 1 @ wly0)
(ehrnat_y1 : wry1 @ 1 = 1 @ wly1)
(ehrnat_z0 : wrz0 @ 1 = 1 @ wlz0)
(ehrnat_z1 : wrz1 @ 1 = 1 @ wlz1)
(ehrnat_yz0 : wryz0 @ 1 = 1 @ wlyz0)
(ehrnat_yz1 : wryz1 @ 1 = 1 @ wlyz1)
(wlrnat_x_y : wlx0 @ wry0 = wry1 @ wlx1)
(wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0)
(wlrnat_x_z : wlx1 @ wrz0 = wrz1 @ wlx2)
(wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0)
(wlrnat_x_yz : wlx0 @ wryz0 = wryz1 @ wlx2)
(wlrnat_yz_x : wlyz1 @ wrx2 = wrx0 @ wlyz0)
(wrpp_yz0 : wry0 @ wrz0 = wryz0)
(wlpp_yz0 : wly0 @ wlz0 = wlyz0)
(wrpp_yz1 : wry1 @ wrz1 = wryz1)
(wlpp_yz1 : wly1 @ wlz1 = wlyz1)
(H_ulnat_yz0 : ulnat_y0 [-] ulnat_z0 =
whiskerR wlpp_yz0 1 @ ulnat_yz0)
(H_urnat_yz0 : urnat_y0 [-] urnat_z0 =
whiskerR wrpp_yz0 1 @ urnat_yz0)
(H_ulnat_yz1 : ulnat_y1 [-] ulnat_z1 =
whiskerR wlpp_yz1 1 @ ulnat_yz1)
(H_urnat_yz1 : urnat_y1 [-] urnat_z1 =
whiskerR wrpp_yz1 1 @ urnat_yz1)
(H_ehrnat_yz0 : (ehrnat_y0 [-] ehrnat_z0) @
whiskerL 1 wlpp_yz0 =
whiskerR wrpp_yz0 1 @ ehrnat_yz0)
(H_ehrnat_yz1 : (ehrnat_y1 [-] ehrnat_z1) @
whiskerL 1 wlpp_yz1 =
whiskerR wrpp_yz1 1 @ ehrnat_yz1)
(H_wlrnat_x_yz : (wlrnat_x_y [I] wlrnat_x_z) @
whiskerR wrpp_yz1 wlx2 =
whiskerL wlx0 wrpp_yz0 @ wlrnat_x_yz)
(H_wlrnat_yz_x : (wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 wlpp_yz0 =
whiskerR wlpp_yz1 wrx2 @ wlrnat_yz_x)
(ehlnat_1p_x0 : (ehlnat_x0 [I] urnat_x0) @ 1 =
1 @ ulnat_x0)
(ehlnat_1p_x1 : (ehlnat_x1 [I] urnat_x1) @ 1 =
1 @ ulnat_x1)
(ehlnat_1p_x2 : (ehlnat_x2 [I] urnat_x2) @ 1 =
1 @ ulnat_x2)
(ehrnat_p1_y0 : (ehrnat_y0 [I] ulnat_y0) @ 1 =
1 @ urnat_y0)
(ehrnat_p1_y1 : (ehrnat_y1 [I] ulnat_y1) @ 1 =
1 @ urnat_y1)
(ehrnat_p1_z0 : (ehrnat_z0 [I] ulnat_z0) @ 1 =
1 @ urnat_z0)
(ehrnat_p1_z1 : (ehrnat_z1 [I] ulnat_z1) @ 1 =
1 @ urnat_z1)
(wlrnat_V_x_y : whiskerR wlrnat_x_y 1 @
(ehrnat_y1 [-] ehlnat_x1) =
(ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^)
(wlrnat_V_x_z : whiskerR wlrnat_x_z 1 @
(ehrnat_z1 [-] ehlnat_x2) =
(ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^),
let eh_x_y :=
concat_p_pp x0 y0 z0 @
whiskerR
(((rlucancel_inv (ulnat_x0 [-] urnat_y0))^ @
wlrnat_x_y) @
rlucancel_inv (urnat_y1 [-] ulnat_x1)) z0 in
(((whiskerR
(concat_p1
(((rlucancel_inv (ulnat_x0 [-] urnat_yz0))^ @
wlrnat_x_yz) @
rlucancel_inv (urnat_yz1 [-] ulnat_x2)) @@
concat_p1
(((rlucancel_inv (ulnat_yz1 [-] urnat_x2))^ @
wlrnat_yz_x) @
rlucancel_inv (urnat_x0 [-] ulnat_yz0)))
eh_x_y @
whiskerR
(eh_V_gen ehlnat_1p_x0 ehlnat_1p_x2
(Ehrnat_p1_pp 1 1 H_ehrnat_yz0 H_ulnat_yz0
H_urnat_yz0 ehrnat_p1_y0 ehrnat_p1_z0)
(Ehrnat_p1_pp 1 1 H_ehrnat_yz1 H_ulnat_yz1
H_urnat_yz1 ehrnat_p1_y1 ehrnat_p1_z1)
(Wlrnat_V_p_pp H_ehrnat_yz0 H_ehrnat_yz1
H_wlrnat_x_yz H_wlrnat_yz_x wlrnat_V_x_y
wlrnat_V_x_z)) eh_x_y) @
lrucancel 1 ) @
whiskerL eh_x_y (concat_pp_p_p_pp y1 x1 z0)^) @
whiskerL eh_x_y
(concat_p1 (concat_pp_p y1 x1 z0) @@
concat_p1 (concat_p_pp y1 x1 z0))^ =
(eh_p_pp_gen H_urnat_yz0 H_urnat_yz1 H_wlrnat_x_yz
[-] lrucancel
(whiskerL (concat_pp_p y1 z1 x2)
(ap
(fun p : z1 @ x2 = x1 @ z0 =>
whiskerL y1 p)
(moveL_V1
(((rlucancel_inv
(ulnat_z1 [-] urnat_x2))^ @
wlrnat_z_x) @
rlucancel_inv (urnat_x1 [-] ulnat_z0))
(((rlucancel_inv
(ulnat_x1 [-] urnat_z0))^ @
wlrnat_x_z) @
rlucancel_inv (urnat_z1 [-] ulnat_x2))
(eh_V_gen ehlnat_1p_x1 ehlnat_1p_x2
ehrnat_p1_z0 ehrnat_p1_z1
wlrnat_V_x_z)))))
[-] (eh_pp_p_gen H_ulnat_yz1 H_ulnat_yz0 H_wlrnat_yz_x
[-] lrucancel
(whiskerL (concat_p_pp x0 y0 z0)
(ap
(fun p : x0 @ y0 = y1 @ x1 =>
whiskerR p z0)
(moveL_1V
(((rlucancel_inv
(ulnat_x0 [-] urnat_y0))^ @
wlrnat_x_y) @
rlucancel_inv
(urnat_y1 [-] ulnat_x1))
(((rlucancel_inv
(ulnat_y1 [-] urnat_x1))^ @
wlrnat_y_x) @
rlucancel_inv
(urnat_x0 [-] ulnat_y0))
(eh_V_gen ehlnat_1p_x0
ehlnat_1p_x1 ehrnat_p1_y0
ehrnat_p1_y1 wlrnat_V_x_y)))))
generalize_goal lem. X : Type a : X
forall
P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3)
(x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0)
(x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2)
(x25 x26 : x = x3) (x27 x28 : a = x2)
(x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5)
(x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8)
(x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6)
(x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26)
(x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28)
(x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type ,
?Goal ->
forall (b c d e f : X) (wlx0 x0 wrx0 : a = b)
(wlx1 x1 wrx1 : c = d) (wlx2 x2 wrx2 : e = f)
(wly0 y0 wry0 : b = d) (wly1 y1 wry1 : a = c)
(wlz0 z0 wrz0 : d = f) (wlz1 z1 wrz1 : c = e)
(wlyz0 wryz0 : b = f) (wlyz1 wryz1 : a = e)
(ulnat_x0 : wlx0 @ 1 = 1 @ x0)
(urnat_x0 : wrx0 @ 1 = 1 @ x0)
(ulnat_x1 : wlx1 @ 1 = 1 @ x1)
(urnat_x1 : wrx1 @ 1 = 1 @ x1)
(ulnat_x2 : wlx2 @ 1 = 1 @ x2)
(urnat_x2 : wrx2 @ 1 = 1 @ x2)
(ulnat_y0 : wly0 @ 1 = 1 @ y0)
(urnat_y0 : wry0 @ 1 = 1 @ y0)
(ulnat_y1 : wly1 @ 1 = 1 @ y1)
(urnat_y1 : wry1 @ 1 = 1 @ y1)
(ulnat_z0 : wlz0 @ 1 = 1 @ z0)
(urnat_z0 : wrz0 @ 1 = 1 @ z0)
(ulnat_z1 : wlz1 @ 1 = 1 @ z1)
(urnat_z1 : wrz1 @ 1 = 1 @ z1)
(ulnat_yz0 : wlyz0 @ 1 = 1 @ (y0 @ z0))
(urnat_yz0 : wryz0 @ 1 = 1 @ (y0 @ z0))
(ulnat_yz1 : wlyz1 @ 1 = 1 @ (y1 @ z1))
(urnat_yz1 : wryz1 @ 1 = 1 @ (y1 @ z1))
(ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0)
(ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1)
(ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2)
(ehrnat_y0 : wry0 @ 1 = 1 @ wly0)
(ehrnat_y1 : wry1 @ 1 = 1 @ wly1)
(ehrnat_z0 : wrz0 @ 1 = 1 @ wlz0)
(ehrnat_z1 : wrz1 @ 1 = 1 @ wlz1)
(ehrnat_yz0 : wryz0 @ 1 = 1 @ wlyz0)
(ehrnat_yz1 : wryz1 @ 1 = 1 @ wlyz1)
(wlrnat_x_y : wlx0 @ wry0 = wry1 @ wlx1)
(wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0)
(wlrnat_x_z : wlx1 @ wrz0 = wrz1 @ wlx2)
(wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0)
(wlrnat_x_yz : wlx0 @ wryz0 = wryz1 @ wlx2)
(wlrnat_yz_x : wlyz1 @ wrx2 = wrx0 @ wlyz0)
(wrpp_yz0 : wry0 @ wrz0 = wryz0)
(wlpp_yz0 : wly0 @ wlz0 = wlyz0)
(wrpp_yz1 : wry1 @ wrz1 = wryz1)
(wlpp_yz1 : wly1 @ wlz1 = wlyz1)
(H_ulnat_yz0 : ulnat_y0 [-] ulnat_z0 =
whiskerR wlpp_yz0 1 @ ulnat_yz0)
(H_urnat_yz0 : urnat_y0 [-] urnat_z0 =
whiskerR wrpp_yz0 1 @ urnat_yz0)
(H_ulnat_yz1 : ulnat_y1 [-] ulnat_z1 =
whiskerR wlpp_yz1 1 @ ulnat_yz1)
(H_urnat_yz1 : urnat_y1 [-] urnat_z1 =
whiskerR wrpp_yz1 1 @ urnat_yz1)
(H_ehrnat_yz0 : (ehrnat_y0 [-] ehrnat_z0) @
whiskerL 1 wlpp_yz0 =
whiskerR wrpp_yz0 1 @ ehrnat_yz0)
(H_ehrnat_yz1 : (ehrnat_y1 [-] ehrnat_z1) @
whiskerL 1 wlpp_yz1 =
whiskerR wrpp_yz1 1 @ ehrnat_yz1)
(H_wlrnat_x_yz : (wlrnat_x_y [I] wlrnat_x_z) @
whiskerR wrpp_yz1 wlx2 =
whiskerL wlx0 wrpp_yz0 @ wlrnat_x_yz)
(H_wlrnat_yz_x : (wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 wlpp_yz0 =
whiskerR wlpp_yz1 wrx2 @ wlrnat_yz_x)
(ehlnat_1p_x0 : (ehlnat_x0 [I] urnat_x0) @ 1 =
1 @ ulnat_x0)
(ehlnat_1p_x1 : (ehlnat_x1 [I] urnat_x1) @ 1 =
1 @ ulnat_x1)
(ehlnat_1p_x2 : (ehlnat_x2 [I] urnat_x2) @ 1 =
1 @ ulnat_x2)
(ehrnat_p1_y0 : (ehrnat_y0 [I] ulnat_y0) @ 1 =
1 @ urnat_y0)
(ehrnat_p1_y1 : (ehrnat_y1 [I] ulnat_y1) @ 1 =
1 @ urnat_y1)
(ehrnat_p1_z0 : (ehrnat_z0 [I] ulnat_z0) @ 1 =
1 @ urnat_z0)
(ehrnat_p1_z1 : (ehrnat_z1 [I] ulnat_z1) @ 1 =
1 @ urnat_z1)
(wlrnat_V_x_y : whiskerR wlrnat_x_y 1 @
(ehrnat_y1 [-] ehlnat_x1) =
(ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^)
(wlrnat_V_x_z : whiskerR wlrnat_x_z 1 @
(ehrnat_z1 [-] ehlnat_x2) =
(ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^),
P b c d e f wlx0 x0 wrx0 wlx1 x1 wrx1 wlx2 x2 wrx2
wly0 y0 wry0 wly1 y1 wry1 wlz0 z0 wrz0 wlz1 z1 wrz1
wlyz0 wryz0 wlyz1 wryz1 ulnat_x0 urnat_x0 ulnat_x1
urnat_x1 ulnat_x2 urnat_x2 ulnat_y0 urnat_y0
ulnat_y1 urnat_y1 ulnat_z0 urnat_z0 ulnat_z1
urnat_z1 ulnat_yz0 urnat_yz0 ulnat_yz1 urnat_yz1
ehlnat_x0 ehlnat_x1 ehlnat_x2 ehrnat_y0 ehrnat_y1
ehrnat_z0 ehrnat_z1 ehrnat_yz0 ehrnat_yz1 wlrnat_x_y
wlrnat_y_x wlrnat_x_z wlrnat_z_x wlrnat_x_yz
wlrnat_yz_x wrpp_yz0 wlpp_yz0 wrpp_yz1 wlpp_yz1
H_ulnat_yz0 H_urnat_yz0 H_ulnat_yz1 H_urnat_yz1
H_ehrnat_yz0 H_ehrnat_yz1 H_wlrnat_x_yz
H_wlrnat_yz_x ehlnat_1p_x0 ehlnat_1p_x1 ehlnat_1p_x2
ehrnat_p1_y0 ehrnat_p1_y1 ehrnat_p1_z0 ehrnat_p1_z1
wlrnat_V_x_y wlrnat_V_x_z
{ X : Type a : X
forall
P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3)
(x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0)
(x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2)
(x25 x26 : x = x3) (x27 x28 : a = x2)
(x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5)
(x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8)
(x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6)
(x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26)
(x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28)
(x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type ,
?Goal ->
forall (b c d e f : X) (wlx0 x0 wrx0 : a = b)
(wlx1 x1 wrx1 : c = d) (wlx2 x2 wrx2 : e = f)
(wly0 y0 wry0 : b = d) (wly1 y1 wry1 : a = c)
(wlz0 z0 wrz0 : d = f) (wlz1 z1 wrz1 : c = e)
(wlyz0 wryz0 : b = f) (wlyz1 wryz1 : a = e)
(ulnat_x0 : wlx0 @ 1 = 1 @ x0)
(urnat_x0 : wrx0 @ 1 = 1 @ x0)
(ulnat_x1 : wlx1 @ 1 = 1 @ x1)
(urnat_x1 : wrx1 @ 1 = 1 @ x1)
(ulnat_x2 : wlx2 @ 1 = 1 @ x2)
(urnat_x2 : wrx2 @ 1 = 1 @ x2)
(ulnat_y0 : wly0 @ 1 = 1 @ y0)
(urnat_y0 : wry0 @ 1 = 1 @ y0)
(ulnat_y1 : wly1 @ 1 = 1 @ y1)
(urnat_y1 : wry1 @ 1 = 1 @ y1)
(ulnat_z0 : wlz0 @ 1 = 1 @ z0)
(urnat_z0 : wrz0 @ 1 = 1 @ z0)
(ulnat_z1 : wlz1 @ 1 = 1 @ z1)
(urnat_z1 : wrz1 @ 1 = 1 @ z1)
(ulnat_yz0 : wlyz0 @ 1 = 1 @ (y0 @ z0))
(urnat_yz0 : wryz0 @ 1 = 1 @ (y0 @ z0))
(ulnat_yz1 : wlyz1 @ 1 = 1 @ (y1 @ z1))
(urnat_yz1 : wryz1 @ 1 = 1 @ (y1 @ z1))
(ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0)
(ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1)
(ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2)
(ehrnat_y0 : wry0 @ 1 = 1 @ wly0)
(ehrnat_y1 : wry1 @ 1 = 1 @ wly1)
(ehrnat_z0 : wrz0 @ 1 = 1 @ wlz0)
(ehrnat_z1 : wrz1 @ 1 = 1 @ wlz1)
(ehrnat_yz0 : wryz0 @ 1 = 1 @ wlyz0)
(ehrnat_yz1 : wryz1 @ 1 = 1 @ wlyz1)
(wlrnat_x_y : wlx0 @ wry0 = wry1 @ wlx1)
(wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0)
(wlrnat_x_z : wlx1 @ wrz0 = wrz1 @ wlx2)
(wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0)
(wlrnat_x_yz : wlx0 @ wryz0 = wryz1 @ wlx2)
(wlrnat_yz_x : wlyz1 @ wrx2 = wrx0 @ wlyz0)
(wrpp_yz0 : wry0 @ wrz0 = wryz0)
(wlpp_yz0 : wly0 @ wlz0 = wlyz0)
(wrpp_yz1 : wry1 @ wrz1 = wryz1)
(wlpp_yz1 : wly1 @ wlz1 = wlyz1)
(H_ulnat_yz0 : ulnat_y0 [-] ulnat_z0 =
whiskerR wlpp_yz0 1 @ ulnat_yz0)
(H_urnat_yz0 : urnat_y0 [-] urnat_z0 =
whiskerR wrpp_yz0 1 @ urnat_yz0)
(H_ulnat_yz1 : ulnat_y1 [-] ulnat_z1 =
whiskerR wlpp_yz1 1 @ ulnat_yz1)
(H_urnat_yz1 : urnat_y1 [-] urnat_z1 =
whiskerR wrpp_yz1 1 @ urnat_yz1)
(H_ehrnat_yz0 : (ehrnat_y0 [-] ehrnat_z0) @
whiskerL 1 wlpp_yz0 =
whiskerR wrpp_yz0 1 @ ehrnat_yz0)
(H_ehrnat_yz1 : (ehrnat_y1 [-] ehrnat_z1) @
whiskerL 1 wlpp_yz1 =
whiskerR wrpp_yz1 1 @ ehrnat_yz1)
(H_wlrnat_x_yz : (wlrnat_x_y [I] wlrnat_x_z) @
whiskerR wrpp_yz1 wlx2 =
whiskerL wlx0 wrpp_yz0 @ wlrnat_x_yz)
(H_wlrnat_yz_x : (wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 wlpp_yz0 =
whiskerR wlpp_yz1 wrx2 @ wlrnat_yz_x)
(ehlnat_1p_x0 : (ehlnat_x0 [I] urnat_x0) @ 1 =
1 @ ulnat_x0)
(ehlnat_1p_x1 : (ehlnat_x1 [I] urnat_x1) @ 1 =
1 @ ulnat_x1)
(ehlnat_1p_x2 : (ehlnat_x2 [I] urnat_x2) @ 1 =
1 @ ulnat_x2)
(ehrnat_p1_y0 : (ehrnat_y0 [I] ulnat_y0) @ 1 =
1 @ urnat_y0)
(ehrnat_p1_y1 : (ehrnat_y1 [I] ulnat_y1) @ 1 =
1 @ urnat_y1)
(ehrnat_p1_z0 : (ehrnat_z0 [I] ulnat_z0) @ 1 =
1 @ urnat_z0)
(ehrnat_p1_z1 : (ehrnat_z1 [I] ulnat_z1) @ 1 =
1 @ urnat_z1)
(wlrnat_V_x_y : whiskerR wlrnat_x_y 1 @
(ehrnat_y1 [-] ehlnat_x1) =
(ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^)
(wlrnat_V_x_z : whiskerR wlrnat_x_z 1 @
(ehrnat_z1 [-] ehlnat_x2) =
(ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^),
P b c d e f wlx0 x0 wrx0 wlx1 x1 wrx1 wlx2 x2 wrx2
wly0 y0 wry0 wly1 y1 wry1 wlz0 z0 wrz0 wlz1 z1 wrz1
wlyz0 wryz0 wlyz1 wryz1 ulnat_x0 urnat_x0 ulnat_x1
urnat_x1 ulnat_x2 urnat_x2 ulnat_y0 urnat_y0
ulnat_y1 urnat_y1 ulnat_z0 urnat_z0 ulnat_z1
urnat_z1 ulnat_yz0 urnat_yz0 ulnat_yz1 urnat_yz1
ehlnat_x0 ehlnat_x1 ehlnat_x2 ehrnat_y0 ehrnat_y1
ehrnat_z0 ehrnat_z1 ehrnat_yz0 ehrnat_yz1 wlrnat_x_y
wlrnat_y_x wlrnat_x_z wlrnat_z_x wlrnat_x_yz
wlrnat_yz_x wrpp_yz0 wlpp_yz0 wrpp_yz1 wlpp_yz1
H_ulnat_yz0 H_urnat_yz0 H_ulnat_yz1 H_urnat_yz1
H_ehrnat_yz0 H_ehrnat_yz1 H_wlrnat_x_yz
H_wlrnat_yz_x ehlnat_1p_x0 ehlnat_1p_x1 ehlnat_1p_x2
ehrnat_p1_y0 ehrnat_p1_y1 ehrnat_p1_z0 ehrnat_p1_z1
wlrnat_V_x_y wlrnat_V_x_z
intros P H; intros .X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b, c, d, e, f : X wlx0, x0, wrx0 : a = b wlx1, x1, wrx1 : c = d wlx2, x2, wrx2 : e = f wly0, y0, wry0 : b = d wly1, y1, wry1 : a = c wlz0, z0, wrz0 : d = f wlz1, z1, wrz1 : c = e wlyz0, wryz0 : b = f wlyz1, wryz1 : a = e ulnat_x0 : wlx0 @ 1 = 1 @ x0 urnat_x0 : wrx0 @ 1 = 1 @ x0 ulnat_x1 : wlx1 @ 1 = 1 @ x1 urnat_x1 : wrx1 @ 1 = 1 @ x1 ulnat_x2 : wlx2 @ 1 = 1 @ x2 urnat_x2 : wrx2 @ 1 = 1 @ x2 ulnat_y0 : wly0 @ 1 = 1 @ y0 urnat_y0 : wry0 @ 1 = 1 @ y0 ulnat_y1 : wly1 @ 1 = 1 @ y1 urnat_y1 : wry1 @ 1 = 1 @ y1 ulnat_z0 : wlz0 @ 1 = 1 @ z0 urnat_z0 : wrz0 @ 1 = 1 @ z0 ulnat_z1 : wlz1 @ 1 = 1 @ z1 urnat_z1 : wrz1 @ 1 = 1 @ z1 ulnat_yz0 : wlyz0 @ 1 = 1 @ (y0 @ z0) urnat_yz0 : wryz0 @ 1 = 1 @ (y0 @ z0) ulnat_yz1 : wlyz1 @ 1 = 1 @ (y1 @ z1) urnat_yz1 : wryz1 @ 1 = 1 @ (y1 @ z1) ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : wry0 @ 1 = 1 @ wly0 ehrnat_y1 : wry1 @ 1 = 1 @ wly1 ehrnat_z0 : wrz0 @ 1 = 1 @ wlz0 ehrnat_z1 : wrz1 @ 1 = 1 @ wlz1 ehrnat_yz0 : wryz0 @ 1 = 1 @ wlyz0 ehrnat_yz1 : wryz1 @ 1 = 1 @ wlyz1 wlrnat_x_y : wlx0 @ wry0 = wry1 @ wlx1 wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_x_z : wlx1 @ wrz0 = wrz1 @ wlx2 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0 wlrnat_x_yz : wlx0 @ wryz0 = wryz1 @ wlx2 wlrnat_yz_x : wlyz1 @ wrx2 = wrx0 @ wlyz0 wrpp_yz0 : wry0 @ wrz0 = wryz0 wlpp_yz0 : wly0 @ wlz0 = wlyz0 wrpp_yz1 : wry1 @ wrz1 = wryz1 wlpp_yz1 : wly1 @ wlz1 = wlyz1 H_ulnat_yz0 : ulnat_y0 [-] ulnat_z0 =
whiskerR wlpp_yz0 1 @ ulnat_yz0 H_urnat_yz0 : urnat_y0 [-] urnat_z0 =
whiskerR wrpp_yz0 1 @ urnat_yz0 H_ulnat_yz1 : ulnat_y1 [-] ulnat_z1 =
whiskerR wlpp_yz1 1 @ ulnat_yz1 H_urnat_yz1 : urnat_y1 [-] urnat_z1 =
whiskerR wrpp_yz1 1 @ urnat_yz1 H_ehrnat_yz0 : (ehrnat_y0 [-] ehrnat_z0) @
whiskerL 1 wlpp_yz0 =
whiskerR wrpp_yz0 1 @ ehrnat_yz0 H_ehrnat_yz1 : (ehrnat_y1 [-] ehrnat_z1) @
whiskerL 1 wlpp_yz1 =
whiskerR wrpp_yz1 1 @ ehrnat_yz1 H_wlrnat_x_yz : (wlrnat_x_y [I] wlrnat_x_z) @
whiskerR wrpp_yz1 wlx2 =
whiskerL wlx0 wrpp_yz0 @ wlrnat_x_yz H_wlrnat_yz_x : (wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 wlpp_yz0 =
whiskerR wlpp_yz1 wrx2 @ wlrnat_yz_x ehlnat_1p_x0 : (ehlnat_x0 [I] urnat_x0) @ 1 =
1 @ ulnat_x0 ehlnat_1p_x1 : (ehlnat_x1 [I] urnat_x1) @ 1 =
1 @ ulnat_x1 ehlnat_1p_x2 : (ehlnat_x2 [I] urnat_x2) @ 1 =
1 @ ulnat_x2 ehrnat_p1_y0 : (ehrnat_y0 [I] ulnat_y0) @ 1 =
1 @ urnat_y0 ehrnat_p1_y1 : (ehrnat_y1 [I] ulnat_y1) @ 1 =
1 @ urnat_y1 ehrnat_p1_z0 : (ehrnat_z0 [I] ulnat_z0) @ 1 =
1 @ urnat_z0 ehrnat_p1_z1 : (ehrnat_z1 [I] ulnat_z1) @ 1 =
1 @ urnat_z1 wlrnat_V_x_y : whiskerR wlrnat_x_y 1 @
(ehrnat_y1 [-] ehlnat_x1) =
(ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^ wlrnat_V_x_z : whiskerR wlrnat_x_z 1 @
(ehrnat_z1 [-] ehlnat_x2) =
(ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^
P b c d e f wlx0 x0 wrx0 wlx1 x1 wrx1 wlx2 x2 wrx2
wly0 y0 wry0 wly1 y1 wry1 wlz0 z0 wrz0 wlz1 z1 wrz1
wlyz0 wryz0 wlyz1 wryz1 ulnat_x0 urnat_x0 ulnat_x1
urnat_x1 ulnat_x2 urnat_x2 ulnat_y0 urnat_y0
ulnat_y1 urnat_y1 ulnat_z0 urnat_z0 ulnat_z1
urnat_z1 ulnat_yz0 urnat_yz0 ulnat_yz1 urnat_yz1
ehlnat_x0 ehlnat_x1 ehlnat_x2 ehrnat_y0 ehrnat_y1
ehrnat_z0 ehrnat_z1 ehrnat_yz0 ehrnat_yz1 wlrnat_x_y
wlrnat_y_x wlrnat_x_z wlrnat_z_x wlrnat_x_yz
wlrnat_yz_x wrpp_yz0 wlpp_yz0 wrpp_yz1 wlpp_yz1
H_ulnat_yz0 H_urnat_yz0 H_ulnat_yz1 H_urnat_yz1
H_ehrnat_yz0 H_ehrnat_yz1 H_wlrnat_x_yz
H_wlrnat_yz_x ehlnat_1p_x0 ehlnat_1p_x1 ehlnat_1p_x2
ehrnat_p1_y0 ehrnat_p1_y1 ehrnat_p1_z0 ehrnat_p1_z1
wlrnat_V_x_y wlrnat_V_x_z
destruct wry0, wry1, wrz0, wrz1.X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, x0, wrx0, wlx1, x1, wrx1, wlx2, x2, wrx2 : a = b wly0, y0 : b = b wly1, y1 : a = a wlz0, z0 : b = b wlz1, z1 : a = a wlyz0, wryz0 : b = b wlyz1, wryz1 : a = a ulnat_x0 : wlx0 @ 1 = 1 @ x0 urnat_x0 : wrx0 @ 1 = 1 @ x0 ulnat_x1 : wlx1 @ 1 = 1 @ x1 urnat_x1 : wrx1 @ 1 = 1 @ x1 ulnat_x2 : wlx2 @ 1 = 1 @ x2 urnat_x2 : wrx2 @ 1 = 1 @ x2 ulnat_y0 : wly0 @ 1 = 1 @ y0 urnat_y0 : 1 @ 1 = 1 @ y0ulnat_y1 : wly1 @ 1 = 1 @ y1 urnat_y1 : 1 @ 1 = 1 @ y1ulnat_z0 : wlz0 @ 1 = 1 @ z0 urnat_z0 : 1 @ 1 = 1 @ z0ulnat_z1 : wlz1 @ 1 = 1 @ z1 urnat_z1 : 1 @ 1 = 1 @ z1ulnat_yz0 : wlyz0 @ 1 = 1 @ (y0 @ z0) urnat_yz0 : wryz0 @ 1 = 1 @ (y0 @ z0) ulnat_yz1 : wlyz1 @ 1 = 1 @ (y1 @ z1) urnat_yz1 : wryz1 @ 1 = 1 @ (y1 @ z1) ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : 1 @ 1 = 1 @ wly0ehrnat_y1 : 1 @ 1 = 1 @ wly1ehrnat_z0 : 1 @ 1 = 1 @ wlz0ehrnat_z1 : 1 @ 1 = 1 @ wlz1ehrnat_yz0 : wryz0 @ 1 = 1 @ wlyz0 ehrnat_yz1 : wryz1 @ 1 = 1 @ wlyz1 wlrnat_x_y : wlx0 @ 1 = 1 @ wlx1 wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_x_z : wlx1 @ 1 = 1 @ wlx2 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0 wlrnat_x_yz : wlx0 @ wryz0 = wryz1 @ wlx2 wlrnat_yz_x : wlyz1 @ wrx2 = wrx0 @ wlyz0 wrpp_yz0 : 1 @ 1 = wryz0wlpp_yz0 : wly0 @ wlz0 = wlyz0 wrpp_yz1 : 1 @ 1 = wryz1wlpp_yz1 : wly1 @ wlz1 = wlyz1 H_ulnat_yz0 : ulnat_y0 [-] ulnat_z0 =
whiskerR wlpp_yz0 1 @ ulnat_yz0 H_urnat_yz0 : urnat_y0 [-] urnat_z0 =
whiskerR wrpp_yz0 1 @ urnat_yz0 H_ulnat_yz1 : ulnat_y1 [-] ulnat_z1 =
whiskerR wlpp_yz1 1 @ ulnat_yz1 H_urnat_yz1 : urnat_y1 [-] urnat_z1 =
whiskerR wrpp_yz1 1 @ urnat_yz1 H_ehrnat_yz0 : (ehrnat_y0 [-] ehrnat_z0) @
whiskerL 1 wlpp_yz0 =
whiskerR wrpp_yz0 1 @ ehrnat_yz0 H_ehrnat_yz1 : (ehrnat_y1 [-] ehrnat_z1) @
whiskerL 1 wlpp_yz1 =
whiskerR wrpp_yz1 1 @ ehrnat_yz1 H_wlrnat_x_yz : (wlrnat_x_y [I] wlrnat_x_z) @
whiskerR wrpp_yz1 wlx2 =
whiskerL wlx0 wrpp_yz0 @ wlrnat_x_yz H_wlrnat_yz_x : (wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 wlpp_yz0 =
whiskerR wlpp_yz1 wrx2 @ wlrnat_yz_x ehlnat_1p_x0 : (ehlnat_x0 [I] urnat_x0) @ 1 =
1 @ ulnat_x0 ehlnat_1p_x1 : (ehlnat_x1 [I] urnat_x1) @ 1 =
1 @ ulnat_x1 ehlnat_1p_x2 : (ehlnat_x2 [I] urnat_x2) @ 1 =
1 @ ulnat_x2 ehrnat_p1_y0 : (ehrnat_y0 [I] ulnat_y0) @ 1 =
1 @ urnat_y0 ehrnat_p1_y1 : (ehrnat_y1 [I] ulnat_y1) @ 1 =
1 @ urnat_y1 ehrnat_p1_z0 : (ehrnat_z0 [I] ulnat_z0) @ 1 =
1 @ urnat_z0 ehrnat_p1_z1 : (ehrnat_z1 [I] ulnat_z1) @ 1 =
1 @ urnat_z1 wlrnat_V_x_y : whiskerR wlrnat_x_y 1 @
(ehrnat_y1 [-] ehlnat_x1) =
(ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^ wlrnat_V_x_z : whiskerR wlrnat_x_z 1 @
(ehrnat_z1 [-] ehlnat_x2) =
(ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^
P b a b a b wlx0 x0 wrx0 wlx1 x1 wrx1 wlx2 x2 wrx2
wly0 y0 1 wly1 y1 1 wlz0 z0 1 wlz1 z1 1 wlyz0 wryz0
wlyz1 wryz1 ulnat_x0 urnat_x0 ulnat_x1 urnat_x1
ulnat_x2 urnat_x2 ulnat_y0 urnat_y0 ulnat_y1
urnat_y1 ulnat_z0 urnat_z0 ulnat_z1 urnat_z1
ulnat_yz0 urnat_yz0 ulnat_yz1 urnat_yz1 ehlnat_x0
ehlnat_x1 ehlnat_x2 ehrnat_y0 ehrnat_y1 ehrnat_z0
ehrnat_z1 ehrnat_yz0 ehrnat_yz1 wlrnat_x_y
wlrnat_y_x wlrnat_x_z wlrnat_z_x wlrnat_x_yz
wlrnat_yz_x wrpp_yz0 wlpp_yz0 wrpp_yz1 wlpp_yz1
H_ulnat_yz0 H_urnat_yz0 H_ulnat_yz1 H_urnat_yz1
H_ehrnat_yz0 H_ehrnat_yz1 H_wlrnat_x_yz
H_wlrnat_yz_x ehlnat_1p_x0 ehlnat_1p_x1 ehlnat_1p_x2
ehrnat_p1_y0 ehrnat_p1_y1 ehrnat_p1_z0 ehrnat_p1_z1
wlrnat_V_x_y wlrnat_V_x_z
destruct wrpp_yz0, wlpp_yz0, wrpp_yz1, wlpp_yz1.X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, x0, wrx0, wlx1, x1, wrx1, wlx2, x2, wrx2 : a = b wly0, y0 : b = b wly1, y1 : a = a wlz0, z0 : b = b wlz1, z1 : a = a ulnat_x0 : wlx0 @ 1 = 1 @ x0 urnat_x0 : wrx0 @ 1 = 1 @ x0 ulnat_x1 : wlx1 @ 1 = 1 @ x1 urnat_x1 : wrx1 @ 1 = 1 @ x1 ulnat_x2 : wlx2 @ 1 = 1 @ x2 urnat_x2 : wrx2 @ 1 = 1 @ x2 ulnat_y0 : wly0 @ 1 = 1 @ y0 urnat_y0 : 1 @ 1 = 1 @ y0ulnat_y1 : wly1 @ 1 = 1 @ y1 urnat_y1 : 1 @ 1 = 1 @ y1ulnat_z0 : wlz0 @ 1 = 1 @ z0 urnat_z0 : 1 @ 1 = 1 @ z0ulnat_z1 : wlz1 @ 1 = 1 @ z1 urnat_z1 : 1 @ 1 = 1 @ z1ulnat_yz0 : (wly0 @ wlz0) @ 1 = 1 @ (y0 @ z0) urnat_yz0 : (1 @ 1 ) @ 1 = 1 @ (y0 @ z0) ulnat_yz1 : (wly1 @ wlz1) @ 1 = 1 @ (y1 @ z1) urnat_yz1 : (1 @ 1 ) @ 1 = 1 @ (y1 @ z1) ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : 1 @ 1 = 1 @ wly0ehrnat_y1 : 1 @ 1 = 1 @ wly1ehrnat_z0 : 1 @ 1 = 1 @ wlz0ehrnat_z1 : 1 @ 1 = 1 @ wlz1ehrnat_yz0 : (1 @ 1 ) @ 1 = 1 @ (wly0 @ wlz0) ehrnat_yz1 : (1 @ 1 ) @ 1 = 1 @ (wly1 @ wlz1) wlrnat_x_y : wlx0 @ 1 = 1 @ wlx1 wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_x_z : wlx1 @ 1 = 1 @ wlx2 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0 wlrnat_x_yz : wlx0 @ (1 @ 1 ) = (1 @ 1 ) @ wlx2 wlrnat_yz_x : (wly1 @ wlz1) @ wrx2 =
wrx0 @ (wly0 @ wlz0) H_ulnat_yz0 : ulnat_y0 [-] ulnat_z0 =
whiskerR 1 1 @ ulnat_yz0 H_urnat_yz0 : urnat_y0 [-] urnat_z0 =
whiskerR 1 1 @ urnat_yz0 H_ulnat_yz1 : ulnat_y1 [-] ulnat_z1 =
whiskerR 1 1 @ ulnat_yz1 H_urnat_yz1 : urnat_y1 [-] urnat_z1 =
whiskerR 1 1 @ urnat_yz1 H_ehrnat_yz0 : (ehrnat_y0 [-] ehrnat_z0) @
whiskerL 1 1 =
whiskerR 1 1 @ ehrnat_yz0 H_ehrnat_yz1 : (ehrnat_y1 [-] ehrnat_z1) @
whiskerL 1 1 =
whiskerR 1 1 @ ehrnat_yz1 H_wlrnat_x_yz : (wlrnat_x_y [I] wlrnat_x_z) @
whiskerR 1 wlx2 =
whiskerL wlx0 1 @ wlrnat_x_yz H_wlrnat_yz_x : (wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 1 =
whiskerR 1 wrx2 @ wlrnat_yz_x ehlnat_1p_x0 : (ehlnat_x0 [I] urnat_x0) @ 1 =
1 @ ulnat_x0 ehlnat_1p_x1 : (ehlnat_x1 [I] urnat_x1) @ 1 =
1 @ ulnat_x1 ehlnat_1p_x2 : (ehlnat_x2 [I] urnat_x2) @ 1 =
1 @ ulnat_x2 ehrnat_p1_y0 : (ehrnat_y0 [I] ulnat_y0) @ 1 =
1 @ urnat_y0 ehrnat_p1_y1 : (ehrnat_y1 [I] ulnat_y1) @ 1 =
1 @ urnat_y1 ehrnat_p1_z0 : (ehrnat_z0 [I] ulnat_z0) @ 1 =
1 @ urnat_z0 ehrnat_p1_z1 : (ehrnat_z1 [I] ulnat_z1) @ 1 =
1 @ urnat_z1 wlrnat_V_x_y : whiskerR wlrnat_x_y 1 @
(ehrnat_y1 [-] ehlnat_x1) =
(ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^ wlrnat_V_x_z : whiskerR wlrnat_x_z 1 @
(ehrnat_z1 [-] ehlnat_x2) =
(ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^
P b a b a b wlx0 x0 wrx0 wlx1 x1 wrx1 wlx2 x2 wrx2
wly0 y0 1 wly1 y1 1 wlz0 z0 1 wlz1 z1 1
(wly0 @ wlz0) (1 @ 1 ) (wly1 @ wlz1)
(1 @ 1 ) ulnat_x0 urnat_x0 ulnat_x1 urnat_x1 ulnat_x2
urnat_x2 ulnat_y0 urnat_y0 ulnat_y1 urnat_y1
ulnat_z0 urnat_z0 ulnat_z1 urnat_z1 ulnat_yz0
urnat_yz0 ulnat_yz1 urnat_yz1 ehlnat_x0 ehlnat_x1
ehlnat_x2 ehrnat_y0 ehrnat_y1 ehrnat_z0 ehrnat_z1
ehrnat_yz0 ehrnat_yz1 wlrnat_x_y wlrnat_y_x
wlrnat_x_z wlrnat_z_x wlrnat_x_yz wlrnat_yz_x 1 1 1
1 H_ulnat_yz0 H_urnat_yz0 H_ulnat_yz1 H_urnat_yz1
H_ehrnat_yz0 H_ehrnat_yz1 H_wlrnat_x_yz
H_wlrnat_yz_x ehlnat_1p_x0 ehlnat_1p_x1 ehlnat_1p_x2
ehrnat_p1_y0 ehrnat_p1_y1 ehrnat_p1_z0 ehrnat_p1_z1
wlrnat_V_x_y wlrnat_V_x_z
revert wlrnat_x_yz H_wlrnat_x_yz.X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, x0, wrx0, wlx1, x1, wrx1, wlx2, x2, wrx2 : a = b wly0, y0 : b = b wly1, y1 : a = a wlz0, z0 : b = b wlz1, z1 : a = a ulnat_x0 : wlx0 @ 1 = 1 @ x0 urnat_x0 : wrx0 @ 1 = 1 @ x0 ulnat_x1 : wlx1 @ 1 = 1 @ x1 urnat_x1 : wrx1 @ 1 = 1 @ x1 ulnat_x2 : wlx2 @ 1 = 1 @ x2 urnat_x2 : wrx2 @ 1 = 1 @ x2 ulnat_y0 : wly0 @ 1 = 1 @ y0 urnat_y0 : 1 @ 1 = 1 @ y0ulnat_y1 : wly1 @ 1 = 1 @ y1 urnat_y1 : 1 @ 1 = 1 @ y1ulnat_z0 : wlz0 @ 1 = 1 @ z0 urnat_z0 : 1 @ 1 = 1 @ z0ulnat_z1 : wlz1 @ 1 = 1 @ z1 urnat_z1 : 1 @ 1 = 1 @ z1ulnat_yz0 : (wly0 @ wlz0) @ 1 = 1 @ (y0 @ z0) urnat_yz0 : (1 @ 1 ) @ 1 = 1 @ (y0 @ z0) ulnat_yz1 : (wly1 @ wlz1) @ 1 = 1 @ (y1 @ z1) urnat_yz1 : (1 @ 1 ) @ 1 = 1 @ (y1 @ z1) ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : 1 @ 1 = 1 @ wly0ehrnat_y1 : 1 @ 1 = 1 @ wly1ehrnat_z0 : 1 @ 1 = 1 @ wlz0ehrnat_z1 : 1 @ 1 = 1 @ wlz1ehrnat_yz0 : (1 @ 1 ) @ 1 = 1 @ (wly0 @ wlz0) ehrnat_yz1 : (1 @ 1 ) @ 1 = 1 @ (wly1 @ wlz1) wlrnat_x_y : wlx0 @ 1 = 1 @ wlx1 wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_x_z : wlx1 @ 1 = 1 @ wlx2 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0 wlrnat_yz_x : (wly1 @ wlz1) @ wrx2 =
wrx0 @ (wly0 @ wlz0) H_ulnat_yz0 : ulnat_y0 [-] ulnat_z0 =
whiskerR 1 1 @ ulnat_yz0 H_urnat_yz0 : urnat_y0 [-] urnat_z0 =
whiskerR 1 1 @ urnat_yz0 H_ulnat_yz1 : ulnat_y1 [-] ulnat_z1 =
whiskerR 1 1 @ ulnat_yz1 H_urnat_yz1 : urnat_y1 [-] urnat_z1 =
whiskerR 1 1 @ urnat_yz1 H_ehrnat_yz0 : (ehrnat_y0 [-] ehrnat_z0) @
whiskerL 1 1 =
whiskerR 1 1 @ ehrnat_yz0 H_ehrnat_yz1 : (ehrnat_y1 [-] ehrnat_z1) @
whiskerL 1 1 =
whiskerR 1 1 @ ehrnat_yz1 H_wlrnat_yz_x : (wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 1 =
whiskerR 1 wrx2 @ wlrnat_yz_x ehlnat_1p_x0 : (ehlnat_x0 [I] urnat_x0) @ 1 =
1 @ ulnat_x0 ehlnat_1p_x1 : (ehlnat_x1 [I] urnat_x1) @ 1 =
1 @ ulnat_x1 ehlnat_1p_x2 : (ehlnat_x2 [I] urnat_x2) @ 1 =
1 @ ulnat_x2 ehrnat_p1_y0 : (ehrnat_y0 [I] ulnat_y0) @ 1 =
1 @ urnat_y0 ehrnat_p1_y1 : (ehrnat_y1 [I] ulnat_y1) @ 1 =
1 @ urnat_y1 ehrnat_p1_z0 : (ehrnat_z0 [I] ulnat_z0) @ 1 =
1 @ urnat_z0 ehrnat_p1_z1 : (ehrnat_z1 [I] ulnat_z1) @ 1 =
1 @ urnat_z1 wlrnat_V_x_y : whiskerR wlrnat_x_y 1 @
(ehrnat_y1 [-] ehlnat_x1) =
(ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^ wlrnat_V_x_z : whiskerR wlrnat_x_z 1 @
(ehrnat_z1 [-] ehlnat_x2) =
(ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^
forall (wlrnat_x_yz : wlx0 @ (1 @ 1 ) = (1 @ 1 ) @ wlx2)
(H_wlrnat_x_yz : (wlrnat_x_y [I] wlrnat_x_z) @
whiskerR 1 wlx2 =
whiskerL wlx0 1 @ wlrnat_x_yz),
P b a b a b wlx0 x0 wrx0 wlx1 x1 wrx1 wlx2 x2 wrx2
wly0 y0 1 wly1 y1 1 wlz0 z0 1 wlz1 z1 1
(wly0 @ wlz0) (1 @ 1 ) (wly1 @ wlz1)
(1 @ 1 ) ulnat_x0 urnat_x0 ulnat_x1 urnat_x1 ulnat_x2
urnat_x2 ulnat_y0 urnat_y0 ulnat_y1 urnat_y1
ulnat_z0 urnat_z0 ulnat_z1 urnat_z1 ulnat_yz0
urnat_yz0 ulnat_yz1 urnat_yz1 ehlnat_x0 ehlnat_x1
ehlnat_x2 ehrnat_y0 ehrnat_y1 ehrnat_z0 ehrnat_z1
ehrnat_yz0 ehrnat_yz1 wlrnat_x_y wlrnat_y_x
wlrnat_x_z wlrnat_z_x wlrnat_x_yz wlrnat_yz_x 1 1 1
1 H_ulnat_yz0 H_urnat_yz0 H_ulnat_yz1 H_urnat_yz1
H_ehrnat_yz0 H_ehrnat_yz1 H_wlrnat_x_yz
H_wlrnat_yz_x ehlnat_1p_x0 ehlnat_1p_x1 ehlnat_1p_x2
ehrnat_p1_y0 ehrnat_p1_y1 ehrnat_p1_z0 ehrnat_p1_z1
wlrnat_V_x_y wlrnat_V_x_z
snapply equiv_path_ind_moveL_Mp. X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, x0, wrx0, wlx1, x1, wrx1, wlx2, x2, wrx2 : a = b wly0, y0 : b = b wly1, y1 : a = a wlz0, z0 : b = b wlz1, z1 : a = a ulnat_x0 : wlx0 @ 1 = 1 @ x0 urnat_x0 : wrx0 @ 1 = 1 @ x0 ulnat_x1 : wlx1 @ 1 = 1 @ x1 urnat_x1 : wrx1 @ 1 = 1 @ x1 ulnat_x2 : wlx2 @ 1 = 1 @ x2 urnat_x2 : wrx2 @ 1 = 1 @ x2 ulnat_y0 : wly0 @ 1 = 1 @ y0 urnat_y0 : 1 @ 1 = 1 @ y0ulnat_y1 : wly1 @ 1 = 1 @ y1 urnat_y1 : 1 @ 1 = 1 @ y1ulnat_z0 : wlz0 @ 1 = 1 @ z0 urnat_z0 : 1 @ 1 = 1 @ z0ulnat_z1 : wlz1 @ 1 = 1 @ z1 urnat_z1 : 1 @ 1 = 1 @ z1ulnat_yz0 : (wly0 @ wlz0) @ 1 = 1 @ (y0 @ z0) urnat_yz0 : (1 @ 1 ) @ 1 = 1 @ (y0 @ z0) ulnat_yz1 : (wly1 @ wlz1) @ 1 = 1 @ (y1 @ z1) urnat_yz1 : (1 @ 1 ) @ 1 = 1 @ (y1 @ z1) ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : 1 @ 1 = 1 @ wly0ehrnat_y1 : 1 @ 1 = 1 @ wly1ehrnat_z0 : 1 @ 1 = 1 @ wlz0ehrnat_z1 : 1 @ 1 = 1 @ wlz1ehrnat_yz0 : (1 @ 1 ) @ 1 = 1 @ (wly0 @ wlz0) ehrnat_yz1 : (1 @ 1 ) @ 1 = 1 @ (wly1 @ wlz1) wlrnat_x_y : wlx0 @ 1 = 1 @ wlx1 wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_x_z : wlx1 @ 1 = 1 @ wlx2 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0 wlrnat_yz_x : (wly1 @ wlz1) @ wrx2 =
wrx0 @ (wly0 @ wlz0) H_ulnat_yz0 : ulnat_y0 [-] ulnat_z0 =
whiskerR 1 1 @ ulnat_yz0 H_urnat_yz0 : urnat_y0 [-] urnat_z0 =
whiskerR 1 1 @ urnat_yz0 H_ulnat_yz1 : ulnat_y1 [-] ulnat_z1 =
whiskerR 1 1 @ ulnat_yz1 H_urnat_yz1 : urnat_y1 [-] urnat_z1 =
whiskerR 1 1 @ urnat_yz1 H_ehrnat_yz0 : (ehrnat_y0 [-] ehrnat_z0) @
whiskerL 1 1 =
whiskerR 1 1 @ ehrnat_yz0 H_ehrnat_yz1 : (ehrnat_y1 [-] ehrnat_z1) @
whiskerL 1 1 =
whiskerR 1 1 @ ehrnat_yz1 H_wlrnat_yz_x : (wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 1 =
whiskerR 1 wrx2 @ wlrnat_yz_x ehlnat_1p_x0 : (ehlnat_x0 [I] urnat_x0) @ 1 =
1 @ ulnat_x0 ehlnat_1p_x1 : (ehlnat_x1 [I] urnat_x1) @ 1 =
1 @ ulnat_x1 ehlnat_1p_x2 : (ehlnat_x2 [I] urnat_x2) @ 1 =
1 @ ulnat_x2 ehrnat_p1_y0 : (ehrnat_y0 [I] ulnat_y0) @ 1 =
1 @ urnat_y0 ehrnat_p1_y1 : (ehrnat_y1 [I] ulnat_y1) @ 1 =
1 @ urnat_y1 ehrnat_p1_z0 : (ehrnat_z0 [I] ulnat_z0) @ 1 =
1 @ urnat_z0 ehrnat_p1_z1 : (ehrnat_z1 [I] ulnat_z1) @ 1 =
1 @ urnat_z1 wlrnat_V_x_y : whiskerR wlrnat_x_y 1 @
(ehrnat_y1 [-] ehlnat_x1) =
(ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^ wlrnat_V_x_z : whiskerR wlrnat_x_z 1 @
(ehrnat_z1 [-] ehlnat_x2) =
(ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^
(fun (q : wlx0 @ (1 @ 1 ) = (1 @ 1 ) @ wlx2)
(s : (wlrnat_x_y [I] wlrnat_x_z) @ whiskerR 1 wlx2 =
whiskerL wlx0 1 @ q) =>
P b a b a b wlx0 x0 wrx0 wlx1 x1 wrx1 wlx2 x2 wrx2
wly0 y0 1 wly1 y1 1 wlz0 z0 1 wlz1 z1 1
(wly0 @ wlz0) (1 @ 1 )
(wly1 @ wlz1) (1 @ 1 ) ulnat_x0 urnat_x0 ulnat_x1
urnat_x1 ulnat_x2 urnat_x2 ulnat_y0 urnat_y0
ulnat_y1 urnat_y1 ulnat_z0 urnat_z0 ulnat_z1
urnat_z1 ulnat_yz0 urnat_yz0 ulnat_yz1 urnat_yz1
ehlnat_x0 ehlnat_x1 ehlnat_x2 ehrnat_y0 ehrnat_y1
ehrnat_z0 ehrnat_z1 ehrnat_yz0 ehrnat_yz1
wlrnat_x_y wlrnat_y_x wlrnat_x_z wlrnat_z_x q
wlrnat_yz_x 1 1 1 1 H_ulnat_yz0 H_urnat_yz0
H_ulnat_yz1 H_urnat_yz1 H_ehrnat_yz0 H_ehrnat_yz1 s
H_wlrnat_yz_x ehlnat_1p_x0 ehlnat_1p_x1
ehlnat_1p_x2 ehrnat_p1_y0 ehrnat_p1_y1 ehrnat_p1_z0
ehrnat_p1_z1 wlrnat_V_x_y wlrnat_V_x_z)
((whiskerL wlx0 1 )^ @
((wlrnat_x_y [I] wlrnat_x_z) @ whiskerR 1 wlx2))
(equiv_moveL_Mp
((whiskerL wlx0 1 )^ @
((wlrnat_x_y [I] wlrnat_x_z) @ whiskerR 1 wlx2))
((wlrnat_x_y [I] wlrnat_x_z) @ whiskerR 1 wlx2)
(whiskerL wlx0 1 ) 1 )
revert wlrnat_x_y wlrnat_V_x_y.X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, x0, wrx0, wlx1, x1, wrx1, wlx2, x2, wrx2 : a = b wly0, y0 : b = b wly1, y1 : a = a wlz0, z0 : b = b wlz1, z1 : a = a ulnat_x0 : wlx0 @ 1 = 1 @ x0 urnat_x0 : wrx0 @ 1 = 1 @ x0 ulnat_x1 : wlx1 @ 1 = 1 @ x1 urnat_x1 : wrx1 @ 1 = 1 @ x1 ulnat_x2 : wlx2 @ 1 = 1 @ x2 urnat_x2 : wrx2 @ 1 = 1 @ x2 ulnat_y0 : wly0 @ 1 = 1 @ y0 urnat_y0 : 1 @ 1 = 1 @ y0ulnat_y1 : wly1 @ 1 = 1 @ y1 urnat_y1 : 1 @ 1 = 1 @ y1ulnat_z0 : wlz0 @ 1 = 1 @ z0 urnat_z0 : 1 @ 1 = 1 @ z0ulnat_z1 : wlz1 @ 1 = 1 @ z1 urnat_z1 : 1 @ 1 = 1 @ z1ulnat_yz0 : (wly0 @ wlz0) @ 1 = 1 @ (y0 @ z0) urnat_yz0 : (1 @ 1 ) @ 1 = 1 @ (y0 @ z0) ulnat_yz1 : (wly1 @ wlz1) @ 1 = 1 @ (y1 @ z1) urnat_yz1 : (1 @ 1 ) @ 1 = 1 @ (y1 @ z1) ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : 1 @ 1 = 1 @ wly0ehrnat_y1 : 1 @ 1 = 1 @ wly1ehrnat_z0 : 1 @ 1 = 1 @ wlz0ehrnat_z1 : 1 @ 1 = 1 @ wlz1ehrnat_yz0 : (1 @ 1 ) @ 1 = 1 @ (wly0 @ wlz0) ehrnat_yz1 : (1 @ 1 ) @ 1 = 1 @ (wly1 @ wlz1) wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_x_z : wlx1 @ 1 = 1 @ wlx2 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0 wlrnat_yz_x : (wly1 @ wlz1) @ wrx2 =
wrx0 @ (wly0 @ wlz0) H_ulnat_yz0 : ulnat_y0 [-] ulnat_z0 =
whiskerR 1 1 @ ulnat_yz0 H_urnat_yz0 : urnat_y0 [-] urnat_z0 =
whiskerR 1 1 @ urnat_yz0 H_ulnat_yz1 : ulnat_y1 [-] ulnat_z1 =
whiskerR 1 1 @ ulnat_yz1 H_urnat_yz1 : urnat_y1 [-] urnat_z1 =
whiskerR 1 1 @ urnat_yz1 H_ehrnat_yz0 : (ehrnat_y0 [-] ehrnat_z0) @
whiskerL 1 1 =
whiskerR 1 1 @ ehrnat_yz0 H_ehrnat_yz1 : (ehrnat_y1 [-] ehrnat_z1) @
whiskerL 1 1 =
whiskerR 1 1 @ ehrnat_yz1 H_wlrnat_yz_x : (wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 1 =
whiskerR 1 wrx2 @ wlrnat_yz_x ehlnat_1p_x0 : (ehlnat_x0 [I] urnat_x0) @ 1 =
1 @ ulnat_x0 ehlnat_1p_x1 : (ehlnat_x1 [I] urnat_x1) @ 1 =
1 @ ulnat_x1 ehlnat_1p_x2 : (ehlnat_x2 [I] urnat_x2) @ 1 =
1 @ ulnat_x2 ehrnat_p1_y0 : (ehrnat_y0 [I] ulnat_y0) @ 1 =
1 @ urnat_y0 ehrnat_p1_y1 : (ehrnat_y1 [I] ulnat_y1) @ 1 =
1 @ urnat_y1 ehrnat_p1_z0 : (ehrnat_z0 [I] ulnat_z0) @ 1 =
1 @ urnat_z0 ehrnat_p1_z1 : (ehrnat_z1 [I] ulnat_z1) @ 1 =
1 @ urnat_z1 wlrnat_V_x_z : whiskerR wlrnat_x_z 1 @
(ehrnat_z1 [-] ehlnat_x2) =
(ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^
forall (wlrnat_x_y : wlx0 @ 1 = 1 @ wlx1)
(wlrnat_V_x_y : whiskerR wlrnat_x_y 1 @
(ehrnat_y1 [-] ehlnat_x1) =
(ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^),
(fun (q : wlx0 @ (1 @ 1 ) = (1 @ 1 ) @ wlx2)
(s : (wlrnat_x_y [I] wlrnat_x_z) @ whiskerR 1 wlx2 =
whiskerL wlx0 1 @ q) =>
P b a b a b wlx0 x0 wrx0 wlx1 x1 wrx1 wlx2 x2 wrx2
wly0 y0 1 wly1 y1 1 wlz0 z0 1 wlz1 z1 1
(wly0 @ wlz0) (1 @ 1 )
(wly1 @ wlz1) (1 @ 1 ) ulnat_x0 urnat_x0 ulnat_x1
urnat_x1 ulnat_x2 urnat_x2 ulnat_y0 urnat_y0
ulnat_y1 urnat_y1 ulnat_z0 urnat_z0 ulnat_z1
urnat_z1 ulnat_yz0 urnat_yz0 ulnat_yz1 urnat_yz1
ehlnat_x0 ehlnat_x1 ehlnat_x2 ehrnat_y0 ehrnat_y1
ehrnat_z0 ehrnat_z1 ehrnat_yz0 ehrnat_yz1
wlrnat_x_y wlrnat_y_x wlrnat_x_z wlrnat_z_x q
wlrnat_yz_x 1 1 1 1 H_ulnat_yz0 H_urnat_yz0
H_ulnat_yz1 H_urnat_yz1 H_ehrnat_yz0 H_ehrnat_yz1 s
H_wlrnat_yz_x ehlnat_1p_x0 ehlnat_1p_x1
ehlnat_1p_x2 ehrnat_p1_y0 ehrnat_p1_y1 ehrnat_p1_z0
ehrnat_p1_z1 wlrnat_V_x_y wlrnat_V_x_z)
((whiskerL wlx0 1 )^ @
((wlrnat_x_y [I] wlrnat_x_z) @ whiskerR 1 wlx2))
(equiv_moveL_Mp
((whiskerL wlx0 1 )^ @
((wlrnat_x_y [I] wlrnat_x_z) @ whiskerR 1 wlx2))
((wlrnat_x_y [I] wlrnat_x_z) @ whiskerR 1 wlx2)
(whiskerL wlx0 1 ) 1 )
snapply (equiv_path_ind (equiv_helper _ _)). X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, x0, wrx0, wlx1, x1, wrx1, wlx2, x2, wrx2 : a = b wly0, y0 : b = b wly1, y1 : a = a wlz0, z0 : b = b wlz1, z1 : a = a ulnat_x0 : wlx0 @ 1 = 1 @ x0 urnat_x0 : wrx0 @ 1 = 1 @ x0 ulnat_x1 : wlx1 @ 1 = 1 @ x1 urnat_x1 : wrx1 @ 1 = 1 @ x1 ulnat_x2 : wlx2 @ 1 = 1 @ x2 urnat_x2 : wrx2 @ 1 = 1 @ x2 ulnat_y0 : wly0 @ 1 = 1 @ y0 urnat_y0 : 1 @ 1 = 1 @ y0ulnat_y1 : wly1 @ 1 = 1 @ y1 urnat_y1 : 1 @ 1 = 1 @ y1ulnat_z0 : wlz0 @ 1 = 1 @ z0 urnat_z0 : 1 @ 1 = 1 @ z0ulnat_z1 : wlz1 @ 1 = 1 @ z1 urnat_z1 : 1 @ 1 = 1 @ z1ulnat_yz0 : (wly0 @ wlz0) @ 1 = 1 @ (y0 @ z0) urnat_yz0 : (1 @ 1 ) @ 1 = 1 @ (y0 @ z0) ulnat_yz1 : (wly1 @ wlz1) @ 1 = 1 @ (y1 @ z1) urnat_yz1 : (1 @ 1 ) @ 1 = 1 @ (y1 @ z1) ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : 1 @ 1 = 1 @ wly0ehrnat_y1 : 1 @ 1 = 1 @ wly1ehrnat_z0 : 1 @ 1 = 1 @ wlz0ehrnat_z1 : 1 @ 1 = 1 @ wlz1ehrnat_yz0 : (1 @ 1 ) @ 1 = 1 @ (wly0 @ wlz0) ehrnat_yz1 : (1 @ 1 ) @ 1 = 1 @ (wly1 @ wlz1) wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_x_z : wlx1 @ 1 = 1 @ wlx2 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0 wlrnat_yz_x : (wly1 @ wlz1) @ wrx2 =
wrx0 @ (wly0 @ wlz0) H_ulnat_yz0 : ulnat_y0 [-] ulnat_z0 =
whiskerR 1 1 @ ulnat_yz0 H_urnat_yz0 : urnat_y0 [-] urnat_z0 =
whiskerR 1 1 @ urnat_yz0 H_ulnat_yz1 : ulnat_y1 [-] ulnat_z1 =
whiskerR 1 1 @ ulnat_yz1 H_urnat_yz1 : urnat_y1 [-] urnat_z1 =
whiskerR 1 1 @ urnat_yz1 H_ehrnat_yz0 : (ehrnat_y0 [-] ehrnat_z0) @
whiskerL 1 1 =
whiskerR 1 1 @ ehrnat_yz0 H_ehrnat_yz1 : (ehrnat_y1 [-] ehrnat_z1) @
whiskerL 1 1 =
whiskerR 1 1 @ ehrnat_yz1 H_wlrnat_yz_x : (wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 1 =
whiskerR 1 wrx2 @ wlrnat_yz_x ehlnat_1p_x0 : (ehlnat_x0 [I] urnat_x0) @ 1 =
1 @ ulnat_x0 ehlnat_1p_x1 : (ehlnat_x1 [I] urnat_x1) @ 1 =
1 @ ulnat_x1 ehlnat_1p_x2 : (ehlnat_x2 [I] urnat_x2) @ 1 =
1 @ ulnat_x2 ehrnat_p1_y0 : (ehrnat_y0 [I] ulnat_y0) @ 1 =
1 @ urnat_y0 ehrnat_p1_y1 : (ehrnat_y1 [I] ulnat_y1) @ 1 =
1 @ urnat_y1 ehrnat_p1_z0 : (ehrnat_z0 [I] ulnat_z0) @ 1 =
1 @ urnat_z0 ehrnat_p1_z1 : (ehrnat_z1 [I] ulnat_z1) @ 1 =
1 @ urnat_z1 wlrnat_V_x_z : whiskerR wlrnat_x_z 1 @
(ehrnat_z1 [-] ehlnat_x2) =
(ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^
(fun (b0 : wlx0 @ 1 = 1 @ wlx1)
(x : (fun s : wlx0 @ 1 = 1 @ wlx1 =>
whiskerR s 1 @ (ehrnat_y1 [-] ehlnat_x1) =
(ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) b0) =>
(fun (q : wlx0 @ (1 @ 1 ) = (1 @ 1 ) @ wlx2)
(s : (b0 [I] wlrnat_x_z) @ whiskerR 1 wlx2 =
whiskerL wlx0 1 @ q) =>
P b a b a b wlx0 x0 wrx0 wlx1 x1 wrx1 wlx2 x2 wrx2
wly0 y0 1 wly1 y1 1 wlz0 z0 1 wlz1 z1 1
(wly0 @ wlz0) (1 @ 1 )
(wly1 @ wlz1) (1 @ 1 ) ulnat_x0 urnat_x0 ulnat_x1
urnat_x1 ulnat_x2 urnat_x2 ulnat_y0 urnat_y0
ulnat_y1 urnat_y1 ulnat_z0 urnat_z0 ulnat_z1
urnat_z1 ulnat_yz0 urnat_yz0 ulnat_yz1 urnat_yz1
ehlnat_x0 ehlnat_x1 ehlnat_x2 ehrnat_y0 ehrnat_y1
ehrnat_z0 ehrnat_z1 ehrnat_yz0 ehrnat_yz1 b0
wlrnat_y_x wlrnat_x_z wlrnat_z_x q wlrnat_yz_x 1 1
1 1 H_ulnat_yz0 H_urnat_yz0 H_ulnat_yz1
H_urnat_yz1 H_ehrnat_yz0 H_ehrnat_yz1 s
H_wlrnat_yz_x ehlnat_1p_x0 ehlnat_1p_x1
ehlnat_1p_x2 ehrnat_p1_y0 ehrnat_p1_y1
ehrnat_p1_z0 ehrnat_p1_z1 x wlrnat_V_x_z)
((whiskerL wlx0 1 )^ @
((b0 [I] wlrnat_x_z) @ whiskerR 1 wlx2))
(equiv_moveL_Mp
((whiskerL wlx0 1 )^ @
((b0 [I] wlrnat_x_z) @ whiskerR 1 wlx2))
((b0 [I] wlrnat_x_z) @ whiskerR 1 wlx2)
(whiskerL wlx0 1 ) 1 ))
(((concat_p1 (wlx0 @ 1 ))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
concat_p1 (1 @ wlx1))
(equiv_helper (ehrnat_y1 [-] ehlnat_x1)
((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^)
(((concat_p1 (wlx0 @ 1 ))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
concat_p1 (1 @ wlx1)) 1 )
revert wlrnat_x_z wlrnat_V_x_z.X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, x0, wrx0, wlx1, x1, wrx1, wlx2, x2, wrx2 : a = b wly0, y0 : b = b wly1, y1 : a = a wlz0, z0 : b = b wlz1, z1 : a = a ulnat_x0 : wlx0 @ 1 = 1 @ x0 urnat_x0 : wrx0 @ 1 = 1 @ x0 ulnat_x1 : wlx1 @ 1 = 1 @ x1 urnat_x1 : wrx1 @ 1 = 1 @ x1 ulnat_x2 : wlx2 @ 1 = 1 @ x2 urnat_x2 : wrx2 @ 1 = 1 @ x2 ulnat_y0 : wly0 @ 1 = 1 @ y0 urnat_y0 : 1 @ 1 = 1 @ y0ulnat_y1 : wly1 @ 1 = 1 @ y1 urnat_y1 : 1 @ 1 = 1 @ y1ulnat_z0 : wlz0 @ 1 = 1 @ z0 urnat_z0 : 1 @ 1 = 1 @ z0ulnat_z1 : wlz1 @ 1 = 1 @ z1 urnat_z1 : 1 @ 1 = 1 @ z1ulnat_yz0 : (wly0 @ wlz0) @ 1 = 1 @ (y0 @ z0) urnat_yz0 : (1 @ 1 ) @ 1 = 1 @ (y0 @ z0) ulnat_yz1 : (wly1 @ wlz1) @ 1 = 1 @ (y1 @ z1) urnat_yz1 : (1 @ 1 ) @ 1 = 1 @ (y1 @ z1) ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : 1 @ 1 = 1 @ wly0ehrnat_y1 : 1 @ 1 = 1 @ wly1ehrnat_z0 : 1 @ 1 = 1 @ wlz0ehrnat_z1 : 1 @ 1 = 1 @ wlz1ehrnat_yz0 : (1 @ 1 ) @ 1 = 1 @ (wly0 @ wlz0) ehrnat_yz1 : (1 @ 1 ) @ 1 = 1 @ (wly1 @ wlz1) wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0 wlrnat_yz_x : (wly1 @ wlz1) @ wrx2 =
wrx0 @ (wly0 @ wlz0) H_ulnat_yz0 : ulnat_y0 [-] ulnat_z0 =
whiskerR 1 1 @ ulnat_yz0 H_urnat_yz0 : urnat_y0 [-] urnat_z0 =
whiskerR 1 1 @ urnat_yz0 H_ulnat_yz1 : ulnat_y1 [-] ulnat_z1 =
whiskerR 1 1 @ ulnat_yz1 H_urnat_yz1 : urnat_y1 [-] urnat_z1 =
whiskerR 1 1 @ urnat_yz1 H_ehrnat_yz0 : (ehrnat_y0 [-] ehrnat_z0) @
whiskerL 1 1 =
whiskerR 1 1 @ ehrnat_yz0 H_ehrnat_yz1 : (ehrnat_y1 [-] ehrnat_z1) @
whiskerL 1 1 =
whiskerR 1 1 @ ehrnat_yz1 H_wlrnat_yz_x : (wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 1 =
whiskerR 1 wrx2 @ wlrnat_yz_x ehlnat_1p_x0 : (ehlnat_x0 [I] urnat_x0) @ 1 =
1 @ ulnat_x0 ehlnat_1p_x1 : (ehlnat_x1 [I] urnat_x1) @ 1 =
1 @ ulnat_x1 ehlnat_1p_x2 : (ehlnat_x2 [I] urnat_x2) @ 1 =
1 @ ulnat_x2 ehrnat_p1_y0 : (ehrnat_y0 [I] ulnat_y0) @ 1 =
1 @ urnat_y0 ehrnat_p1_y1 : (ehrnat_y1 [I] ulnat_y1) @ 1 =
1 @ urnat_y1 ehrnat_p1_z0 : (ehrnat_z0 [I] ulnat_z0) @ 1 =
1 @ urnat_z0 ehrnat_p1_z1 : (ehrnat_z1 [I] ulnat_z1) @ 1 =
1 @ urnat_z1
forall (wlrnat_x_z : wlx1 @ 1 = 1 @ wlx2)
(wlrnat_V_x_z : whiskerR wlrnat_x_z 1 @
(ehrnat_z1 [-] ehlnat_x2) =
(ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^),
(fun (b0 : wlx0 @ 1 = 1 @ wlx1)
(x : (fun s : wlx0 @ 1 = 1 @ wlx1 =>
whiskerR s 1 @ (ehrnat_y1 [-] ehlnat_x1) =
(ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) b0) =>
(fun (q : wlx0 @ (1 @ 1 ) = (1 @ 1 ) @ wlx2)
(s : (b0 [I] wlrnat_x_z) @ whiskerR 1 wlx2 =
whiskerL wlx0 1 @ q) =>
P b a b a b wlx0 x0 wrx0 wlx1 x1 wrx1 wlx2 x2 wrx2
wly0 y0 1 wly1 y1 1 wlz0 z0 1 wlz1 z1 1
(wly0 @ wlz0) (1 @ 1 )
(wly1 @ wlz1) (1 @ 1 ) ulnat_x0 urnat_x0 ulnat_x1
urnat_x1 ulnat_x2 urnat_x2 ulnat_y0 urnat_y0
ulnat_y1 urnat_y1 ulnat_z0 urnat_z0 ulnat_z1
urnat_z1 ulnat_yz0 urnat_yz0 ulnat_yz1 urnat_yz1
ehlnat_x0 ehlnat_x1 ehlnat_x2 ehrnat_y0 ehrnat_y1
ehrnat_z0 ehrnat_z1 ehrnat_yz0 ehrnat_yz1 b0
wlrnat_y_x wlrnat_x_z wlrnat_z_x q wlrnat_yz_x 1 1
1 1 H_ulnat_yz0 H_urnat_yz0 H_ulnat_yz1
H_urnat_yz1 H_ehrnat_yz0 H_ehrnat_yz1 s
H_wlrnat_yz_x ehlnat_1p_x0 ehlnat_1p_x1
ehlnat_1p_x2 ehrnat_p1_y0 ehrnat_p1_y1
ehrnat_p1_z0 ehrnat_p1_z1 x wlrnat_V_x_z)
((whiskerL wlx0 1 )^ @
((b0 [I] wlrnat_x_z) @ whiskerR 1 wlx2))
(equiv_moveL_Mp
((whiskerL wlx0 1 )^ @
((b0 [I] wlrnat_x_z) @ whiskerR 1 wlx2))
((b0 [I] wlrnat_x_z) @ whiskerR 1 wlx2)
(whiskerL wlx0 1 ) 1 ))
(((concat_p1 (wlx0 @ 1 ))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
concat_p1 (1 @ wlx1))
(equiv_helper (ehrnat_y1 [-] ehlnat_x1)
((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^)
(((concat_p1 (wlx0 @ 1 ))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
concat_p1 (1 @ wlx1)) 1 )
snapply (equiv_path_ind (equiv_helper _ _)). X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, x0, wrx0, wlx1, x1, wrx1, wlx2, x2, wrx2 : a = b wly0, y0 : b = b wly1, y1 : a = a wlz0, z0 : b = b wlz1, z1 : a = a ulnat_x0 : wlx0 @ 1 = 1 @ x0 urnat_x0 : wrx0 @ 1 = 1 @ x0 ulnat_x1 : wlx1 @ 1 = 1 @ x1 urnat_x1 : wrx1 @ 1 = 1 @ x1 ulnat_x2 : wlx2 @ 1 = 1 @ x2 urnat_x2 : wrx2 @ 1 = 1 @ x2 ulnat_y0 : wly0 @ 1 = 1 @ y0 urnat_y0 : 1 @ 1 = 1 @ y0ulnat_y1 : wly1 @ 1 = 1 @ y1 urnat_y1 : 1 @ 1 = 1 @ y1ulnat_z0 : wlz0 @ 1 = 1 @ z0 urnat_z0 : 1 @ 1 = 1 @ z0ulnat_z1 : wlz1 @ 1 = 1 @ z1 urnat_z1 : 1 @ 1 = 1 @ z1ulnat_yz0 : (wly0 @ wlz0) @ 1 = 1 @ (y0 @ z0) urnat_yz0 : (1 @ 1 ) @ 1 = 1 @ (y0 @ z0) ulnat_yz1 : (wly1 @ wlz1) @ 1 = 1 @ (y1 @ z1) urnat_yz1 : (1 @ 1 ) @ 1 = 1 @ (y1 @ z1) ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : 1 @ 1 = 1 @ wly0ehrnat_y1 : 1 @ 1 = 1 @ wly1ehrnat_z0 : 1 @ 1 = 1 @ wlz0ehrnat_z1 : 1 @ 1 = 1 @ wlz1ehrnat_yz0 : (1 @ 1 ) @ 1 = 1 @ (wly0 @ wlz0) ehrnat_yz1 : (1 @ 1 ) @ 1 = 1 @ (wly1 @ wlz1) wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0 wlrnat_yz_x : (wly1 @ wlz1) @ wrx2 =
wrx0 @ (wly0 @ wlz0) H_ulnat_yz0 : ulnat_y0 [-] ulnat_z0 =
whiskerR 1 1 @ ulnat_yz0 H_urnat_yz0 : urnat_y0 [-] urnat_z0 =
whiskerR 1 1 @ urnat_yz0 H_ulnat_yz1 : ulnat_y1 [-] ulnat_z1 =
whiskerR 1 1 @ ulnat_yz1 H_urnat_yz1 : urnat_y1 [-] urnat_z1 =
whiskerR 1 1 @ urnat_yz1 H_ehrnat_yz0 : (ehrnat_y0 [-] ehrnat_z0) @
whiskerL 1 1 =
whiskerR 1 1 @ ehrnat_yz0 H_ehrnat_yz1 : (ehrnat_y1 [-] ehrnat_z1) @
whiskerL 1 1 =
whiskerR 1 1 @ ehrnat_yz1 H_wlrnat_yz_x : (wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 1 =
whiskerR 1 wrx2 @ wlrnat_yz_x ehlnat_1p_x0 : (ehlnat_x0 [I] urnat_x0) @ 1 =
1 @ ulnat_x0 ehlnat_1p_x1 : (ehlnat_x1 [I] urnat_x1) @ 1 =
1 @ ulnat_x1 ehlnat_1p_x2 : (ehlnat_x2 [I] urnat_x2) @ 1 =
1 @ ulnat_x2 ehrnat_p1_y0 : (ehrnat_y0 [I] ulnat_y0) @ 1 =
1 @ urnat_y0 ehrnat_p1_y1 : (ehrnat_y1 [I] ulnat_y1) @ 1 =
1 @ urnat_y1 ehrnat_p1_z0 : (ehrnat_z0 [I] ulnat_z0) @ 1 =
1 @ urnat_z0 ehrnat_p1_z1 : (ehrnat_z1 [I] ulnat_z1) @ 1 =
1 @ urnat_z1
(fun (b0 : wlx1 @ 1 = 1 @ wlx2)
(x : (fun s : wlx1 @ 1 = 1 @ wlx2 =>
whiskerR s 1 @ (ehrnat_z1 [-] ehlnat_x2) =
(ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^) b0) =>
(fun (b1 : wlx0 @ 1 = 1 @ wlx1)
(x3 : (fun s : wlx0 @ 1 = 1 @ wlx1 =>
whiskerR s 1 @ (ehrnat_y1 [-] ehlnat_x1) =
(ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) b1) =>
(fun (q : wlx0 @ (1 @ 1 ) = (1 @ 1 ) @ wlx2)
(s : (b1 [I] b0) @ whiskerR 1 wlx2 =
whiskerL wlx0 1 @ q) =>
P b a b a b wlx0 x0 wrx0 wlx1 x1 wrx1 wlx2 x2 wrx2
wly0 y0 1 wly1 y1 1 wlz0 z0 1 wlz1 z1 1
(wly0 @ wlz0) (1 @ 1 )
(wly1 @ wlz1) (1 @ 1 ) ulnat_x0 urnat_x0 ulnat_x1
urnat_x1 ulnat_x2 urnat_x2 ulnat_y0 urnat_y0
ulnat_y1 urnat_y1 ulnat_z0 urnat_z0 ulnat_z1
urnat_z1 ulnat_yz0 urnat_yz0 ulnat_yz1 urnat_yz1
ehlnat_x0 ehlnat_x1 ehlnat_x2 ehrnat_y0 ehrnat_y1
ehrnat_z0 ehrnat_z1 ehrnat_yz0 ehrnat_yz1 b1
wlrnat_y_x b0 wlrnat_z_x q wlrnat_yz_x 1 1 1 1
H_ulnat_yz0 H_urnat_yz0 H_ulnat_yz1 H_urnat_yz1
H_ehrnat_yz0 H_ehrnat_yz1 s H_wlrnat_yz_x
ehlnat_1p_x0 ehlnat_1p_x1 ehlnat_1p_x2
ehrnat_p1_y0 ehrnat_p1_y1 ehrnat_p1_z0
ehrnat_p1_z1 x3 x)
((whiskerL wlx0 1 )^ @
((b1 [I] b0) @ whiskerR 1 wlx2))
(equiv_moveL_Mp
((whiskerL wlx0 1 )^ @
((b1 [I] b0) @ whiskerR 1 wlx2))
((b1 [I] b0) @ whiskerR 1 wlx2)
(whiskerL wlx0 1 ) 1 ))
(((concat_p1 (wlx0 @ 1 ))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
concat_p1 (1 @ wlx1))
(equiv_helper (ehrnat_y1 [-] ehlnat_x1)
((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^)
(((concat_p1 (wlx0 @ 1 ))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
concat_p1 (1 @ wlx1)) 1 ))
(((concat_p1 (wlx1 @ 1 ))^ @
(((ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^) @
(ehrnat_z1 [-] ehlnat_x2)^)) @
concat_p1 (1 @ wlx2))
(equiv_helper (ehrnat_z1 [-] ehlnat_x2)
((ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^)
(((concat_p1 (wlx1 @ 1 ))^ @
(((ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^) @
(ehrnat_z1 [-] ehlnat_x2)^)) @
concat_p1 (1 @ wlx2)) 1 )
revert ulnat_x0 ehlnat_1p_x0.X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, x0, wrx0, wlx1, x1, wrx1, wlx2, x2, wrx2 : a = b wly0, y0 : b = b wly1, y1 : a = a wlz0, z0 : b = b wlz1, z1 : a = a urnat_x0 : wrx0 @ 1 = 1 @ x0 ulnat_x1 : wlx1 @ 1 = 1 @ x1 urnat_x1 : wrx1 @ 1 = 1 @ x1 ulnat_x2 : wlx2 @ 1 = 1 @ x2 urnat_x2 : wrx2 @ 1 = 1 @ x2 ulnat_y0 : wly0 @ 1 = 1 @ y0 urnat_y0 : 1 @ 1 = 1 @ y0ulnat_y1 : wly1 @ 1 = 1 @ y1 urnat_y1 : 1 @ 1 = 1 @ y1ulnat_z0 : wlz0 @ 1 = 1 @ z0 urnat_z0 : 1 @ 1 = 1 @ z0ulnat_z1 : wlz1 @ 1 = 1 @ z1 urnat_z1 : 1 @ 1 = 1 @ z1ulnat_yz0 : (wly0 @ wlz0) @ 1 = 1 @ (y0 @ z0) urnat_yz0 : (1 @ 1 ) @ 1 = 1 @ (y0 @ z0) ulnat_yz1 : (wly1 @ wlz1) @ 1 = 1 @ (y1 @ z1) urnat_yz1 : (1 @ 1 ) @ 1 = 1 @ (y1 @ z1) ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : 1 @ 1 = 1 @ wly0ehrnat_y1 : 1 @ 1 = 1 @ wly1ehrnat_z0 : 1 @ 1 = 1 @ wlz0ehrnat_z1 : 1 @ 1 = 1 @ wlz1ehrnat_yz0 : (1 @ 1 ) @ 1 = 1 @ (wly0 @ wlz0) ehrnat_yz1 : (1 @ 1 ) @ 1 = 1 @ (wly1 @ wlz1) wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0 wlrnat_yz_x : (wly1 @ wlz1) @ wrx2 =
wrx0 @ (wly0 @ wlz0) H_ulnat_yz0 : ulnat_y0 [-] ulnat_z0 =
whiskerR 1 1 @ ulnat_yz0 H_urnat_yz0 : urnat_y0 [-] urnat_z0 =
whiskerR 1 1 @ urnat_yz0 H_ulnat_yz1 : ulnat_y1 [-] ulnat_z1 =
whiskerR 1 1 @ ulnat_yz1 H_urnat_yz1 : urnat_y1 [-] urnat_z1 =
whiskerR 1 1 @ urnat_yz1 H_ehrnat_yz0 : (ehrnat_y0 [-] ehrnat_z0) @
whiskerL 1 1 =
whiskerR 1 1 @ ehrnat_yz0 H_ehrnat_yz1 : (ehrnat_y1 [-] ehrnat_z1) @
whiskerL 1 1 =
whiskerR 1 1 @ ehrnat_yz1 H_wlrnat_yz_x : (wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 1 =
whiskerR 1 wrx2 @ wlrnat_yz_x ehlnat_1p_x1 : (ehlnat_x1 [I] urnat_x1) @ 1 =
1 @ ulnat_x1 ehlnat_1p_x2 : (ehlnat_x2 [I] urnat_x2) @ 1 =
1 @ ulnat_x2 ehrnat_p1_y0 : (ehrnat_y0 [I] ulnat_y0) @ 1 =
1 @ urnat_y0 ehrnat_p1_y1 : (ehrnat_y1 [I] ulnat_y1) @ 1 =
1 @ urnat_y1 ehrnat_p1_z0 : (ehrnat_z0 [I] ulnat_z0) @ 1 =
1 @ urnat_z0 ehrnat_p1_z1 : (ehrnat_z1 [I] ulnat_z1) @ 1 =
1 @ urnat_z1
forall (ulnat_x0 : wlx0 @ 1 = 1 @ x0)
(ehlnat_1p_x0 : (ehlnat_x0 [I] urnat_x0) @ 1 =
1 @ ulnat_x0),
(fun (b0 : wlx1 @ 1 = 1 @ wlx2)
(x : (fun s : wlx1 @ 1 = 1 @ wlx2 =>
whiskerR s 1 @ (ehrnat_z1 [-] ehlnat_x2) =
(ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^) b0) =>
(fun (b1 : wlx0 @ 1 = 1 @ wlx1)
(x3 : (fun s : wlx0 @ 1 = 1 @ wlx1 =>
whiskerR s 1 @ (ehrnat_y1 [-] ehlnat_x1) =
(ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) b1) =>
(fun (q : wlx0 @ (1 @ 1 ) = (1 @ 1 ) @ wlx2)
(s : (b1 [I] b0) @ whiskerR 1 wlx2 =
whiskerL wlx0 1 @ q) =>
P b a b a b wlx0 x0 wrx0 wlx1 x1 wrx1 wlx2 x2 wrx2
wly0 y0 1 wly1 y1 1 wlz0 z0 1 wlz1 z1 1
(wly0 @ wlz0) (1 @ 1 )
(wly1 @ wlz1) (1 @ 1 ) ulnat_x0 urnat_x0 ulnat_x1
urnat_x1 ulnat_x2 urnat_x2 ulnat_y0 urnat_y0
ulnat_y1 urnat_y1 ulnat_z0 urnat_z0 ulnat_z1
urnat_z1 ulnat_yz0 urnat_yz0 ulnat_yz1 urnat_yz1
ehlnat_x0 ehlnat_x1 ehlnat_x2 ehrnat_y0 ehrnat_y1
ehrnat_z0 ehrnat_z1 ehrnat_yz0 ehrnat_yz1 b1
wlrnat_y_x b0 wlrnat_z_x q wlrnat_yz_x 1 1 1 1
H_ulnat_yz0 H_urnat_yz0 H_ulnat_yz1 H_urnat_yz1
H_ehrnat_yz0 H_ehrnat_yz1 s H_wlrnat_yz_x
ehlnat_1p_x0 ehlnat_1p_x1 ehlnat_1p_x2
ehrnat_p1_y0 ehrnat_p1_y1 ehrnat_p1_z0
ehrnat_p1_z1 x3 x)
((whiskerL wlx0 1 )^ @
((b1 [I] b0) @ whiskerR 1 wlx2))
(equiv_moveL_Mp
((whiskerL wlx0 1 )^ @
((b1 [I] b0) @ whiskerR 1 wlx2))
((b1 [I] b0) @ whiskerR 1 wlx2)
(whiskerL wlx0 1 ) 1 ))
(((concat_p1 (wlx0 @ 1 ))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
concat_p1 (1 @ wlx1))
(equiv_helper (ehrnat_y1 [-] ehlnat_x1)
((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^)
(((concat_p1 (wlx0 @ 1 ))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
concat_p1 (1 @ wlx1)) 1 ))
(((concat_p1 (wlx1 @ 1 ))^ @
(((ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^) @
(ehrnat_z1 [-] ehlnat_x2)^)) @
concat_p1 (1 @ wlx2))
(equiv_helper (ehrnat_z1 [-] ehlnat_x2)
((ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^)
(((concat_p1 (wlx1 @ 1 ))^ @
(((ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^) @
(ehrnat_z1 [-] ehlnat_x2)^)) @
concat_p1 (1 @ wlx2)) 1 )
snapply equiv_path_ind_rlucancel. X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, x0, wrx0, wlx1, x1, wrx1, wlx2, x2, wrx2 : a = b wly0, y0 : b = b wly1, y1 : a = a wlz0, z0 : b = b wlz1, z1 : a = a urnat_x0 : wrx0 @ 1 = 1 @ x0 ulnat_x1 : wlx1 @ 1 = 1 @ x1 urnat_x1 : wrx1 @ 1 = 1 @ x1 ulnat_x2 : wlx2 @ 1 = 1 @ x2 urnat_x2 : wrx2 @ 1 = 1 @ x2 ulnat_y0 : wly0 @ 1 = 1 @ y0 urnat_y0 : 1 @ 1 = 1 @ y0ulnat_y1 : wly1 @ 1 = 1 @ y1 urnat_y1 : 1 @ 1 = 1 @ y1ulnat_z0 : wlz0 @ 1 = 1 @ z0 urnat_z0 : 1 @ 1 = 1 @ z0ulnat_z1 : wlz1 @ 1 = 1 @ z1 urnat_z1 : 1 @ 1 = 1 @ z1ulnat_yz0 : (wly0 @ wlz0) @ 1 = 1 @ (y0 @ z0) urnat_yz0 : (1 @ 1 ) @ 1 = 1 @ (y0 @ z0) ulnat_yz1 : (wly1 @ wlz1) @ 1 = 1 @ (y1 @ z1) urnat_yz1 : (1 @ 1 ) @ 1 = 1 @ (y1 @ z1) ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : 1 @ 1 = 1 @ wly0ehrnat_y1 : 1 @ 1 = 1 @ wly1ehrnat_z0 : 1 @ 1 = 1 @ wlz0ehrnat_z1 : 1 @ 1 = 1 @ wlz1ehrnat_yz0 : (1 @ 1 ) @ 1 = 1 @ (wly0 @ wlz0) ehrnat_yz1 : (1 @ 1 ) @ 1 = 1 @ (wly1 @ wlz1) wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0 wlrnat_yz_x : (wly1 @ wlz1) @ wrx2 =
wrx0 @ (wly0 @ wlz0) H_ulnat_yz0 : ulnat_y0 [-] ulnat_z0 =
whiskerR 1 1 @ ulnat_yz0 H_urnat_yz0 : urnat_y0 [-] urnat_z0 =
whiskerR 1 1 @ urnat_yz0 H_ulnat_yz1 : ulnat_y1 [-] ulnat_z1 =
whiskerR 1 1 @ ulnat_yz1 H_urnat_yz1 : urnat_y1 [-] urnat_z1 =
whiskerR 1 1 @ urnat_yz1 H_ehrnat_yz0 : (ehrnat_y0 [-] ehrnat_z0) @
whiskerL 1 1 =
whiskerR 1 1 @ ehrnat_yz0 H_ehrnat_yz1 : (ehrnat_y1 [-] ehrnat_z1) @
whiskerL 1 1 =
whiskerR 1 1 @ ehrnat_yz1 H_wlrnat_yz_x : (wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 1 =
whiskerR 1 wrx2 @ wlrnat_yz_x ehlnat_1p_x1 : (ehlnat_x1 [I] urnat_x1) @ 1 =
1 @ ulnat_x1 ehlnat_1p_x2 : (ehlnat_x2 [I] urnat_x2) @ 1 =
1 @ ulnat_x2 ehrnat_p1_y0 : (ehrnat_y0 [I] ulnat_y0) @ 1 =
1 @ urnat_y0 ehrnat_p1_y1 : (ehrnat_y1 [I] ulnat_y1) @ 1 =
1 @ urnat_y1 ehrnat_p1_z0 : (ehrnat_z0 [I] ulnat_z0) @ 1 =
1 @ urnat_z0 ehrnat_p1_z1 : (ehrnat_z1 [I] ulnat_z1) @ 1 =
1 @ urnat_z1
(fun (q : wlx0 @ 1 = 1 @ x0)
(s : (ehlnat_x0 [I] urnat_x0) @ 1 = 1 @ q) =>
(fun (b0 : wlx1 @ 1 = 1 @ wlx2)
(x : (fun s0 : wlx1 @ 1 = 1 @ wlx2 =>
whiskerR s0 1 @ (ehrnat_z1 [-] ehlnat_x2) =
(ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^) b0) =>
(fun (b1 : wlx0 @ 1 = 1 @ wlx1)
(x3 : (fun s0 : wlx0 @ 1 = 1 @ wlx1 =>
whiskerR s0 1 @ (ehrnat_y1 [-] ehlnat_x1) =
(ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) b1) =>
(fun (q0 : wlx0 @ (1 @ 1 ) = (1 @ 1 ) @ wlx2)
(s0 : (b1 [I] b0) @ whiskerR 1 wlx2 =
whiskerL wlx0 1 @ q0) =>
P b a b a b wlx0 x0 wrx0 wlx1 x1 wrx1 wlx2 x2 wrx2
wly0 y0 1 wly1 y1 1 wlz0 z0 1 wlz1 z1 1
(wly0 @ wlz0) (1 @ 1 )
(wly1 @ wlz1) (1 @ 1 ) q urnat_x0 ulnat_x1
urnat_x1 ulnat_x2 urnat_x2 ulnat_y0 urnat_y0
ulnat_y1 urnat_y1 ulnat_z0 urnat_z0 ulnat_z1
urnat_z1 ulnat_yz0 urnat_yz0 ulnat_yz1 urnat_yz1
ehlnat_x0 ehlnat_x1 ehlnat_x2 ehrnat_y0
ehrnat_y1 ehrnat_z0 ehrnat_z1 ehrnat_yz0
ehrnat_yz1 b1 wlrnat_y_x b0 wlrnat_z_x q0
wlrnat_yz_x 1 1 1 1 H_ulnat_yz0 H_urnat_yz0
H_ulnat_yz1 H_urnat_yz1 H_ehrnat_yz0
H_ehrnat_yz1 s0 H_wlrnat_yz_x s ehlnat_1p_x1
ehlnat_1p_x2 ehrnat_p1_y0 ehrnat_p1_y1
ehrnat_p1_z0 ehrnat_p1_z1 x3 x)
((whiskerL wlx0 1 )^ @
((b1 [I] b0) @ whiskerR 1 wlx2))
(equiv_moveL_Mp
((whiskerL wlx0 1 )^ @
((b1 [I] b0) @ whiskerR 1 wlx2))
((b1 [I] b0) @ whiskerR 1 wlx2)
(whiskerL wlx0 1 ) 1 ))
(((concat_p1 (wlx0 @ 1 ))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
concat_p1 (1 @ wlx1))
(equiv_helper (ehrnat_y1 [-] ehlnat_x1)
((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^)
(((concat_p1 (wlx0 @ 1 ))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
concat_p1 (1 @ wlx1)) 1 ))
(((concat_p1 (wlx1 @ 1 ))^ @
(((ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^) @
(ehrnat_z1 [-] ehlnat_x2)^)) @
concat_p1 (1 @ wlx2))
(equiv_helper (ehrnat_z1 [-] ehlnat_x2)
((ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^)
(((concat_p1 (wlx1 @ 1 ))^ @
(((ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^) @
(ehrnat_z1 [-] ehlnat_x2)^)) @
concat_p1 (1 @ wlx2)) 1 ))
(ehlnat_x0 [I] urnat_x0)
(rlucancel 1 )
revert ulnat_x1 ehlnat_1p_x1.X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, x0, wrx0, wlx1, x1, wrx1, wlx2, x2, wrx2 : a = b wly0, y0 : b = b wly1, y1 : a = a wlz0, z0 : b = b wlz1, z1 : a = a urnat_x0 : wrx0 @ 1 = 1 @ x0 urnat_x1 : wrx1 @ 1 = 1 @ x1 ulnat_x2 : wlx2 @ 1 = 1 @ x2 urnat_x2 : wrx2 @ 1 = 1 @ x2 ulnat_y0 : wly0 @ 1 = 1 @ y0 urnat_y0 : 1 @ 1 = 1 @ y0ulnat_y1 : wly1 @ 1 = 1 @ y1 urnat_y1 : 1 @ 1 = 1 @ y1ulnat_z0 : wlz0 @ 1 = 1 @ z0 urnat_z0 : 1 @ 1 = 1 @ z0ulnat_z1 : wlz1 @ 1 = 1 @ z1 urnat_z1 : 1 @ 1 = 1 @ z1ulnat_yz0 : (wly0 @ wlz0) @ 1 = 1 @ (y0 @ z0) urnat_yz0 : (1 @ 1 ) @ 1 = 1 @ (y0 @ z0) ulnat_yz1 : (wly1 @ wlz1) @ 1 = 1 @ (y1 @ z1) urnat_yz1 : (1 @ 1 ) @ 1 = 1 @ (y1 @ z1) ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : 1 @ 1 = 1 @ wly0ehrnat_y1 : 1 @ 1 = 1 @ wly1ehrnat_z0 : 1 @ 1 = 1 @ wlz0ehrnat_z1 : 1 @ 1 = 1 @ wlz1ehrnat_yz0 : (1 @ 1 ) @ 1 = 1 @ (wly0 @ wlz0) ehrnat_yz1 : (1 @ 1 ) @ 1 = 1 @ (wly1 @ wlz1) wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0 wlrnat_yz_x : (wly1 @ wlz1) @ wrx2 =
wrx0 @ (wly0 @ wlz0) H_ulnat_yz0 : ulnat_y0 [-] ulnat_z0 =
whiskerR 1 1 @ ulnat_yz0 H_urnat_yz0 : urnat_y0 [-] urnat_z0 =
whiskerR 1 1 @ urnat_yz0 H_ulnat_yz1 : ulnat_y1 [-] ulnat_z1 =
whiskerR 1 1 @ ulnat_yz1 H_urnat_yz1 : urnat_y1 [-] urnat_z1 =
whiskerR 1 1 @ urnat_yz1 H_ehrnat_yz0 : (ehrnat_y0 [-] ehrnat_z0) @
whiskerL 1 1 =
whiskerR 1 1 @ ehrnat_yz0 H_ehrnat_yz1 : (ehrnat_y1 [-] ehrnat_z1) @
whiskerL 1 1 =
whiskerR 1 1 @ ehrnat_yz1 H_wlrnat_yz_x : (wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 1 =
whiskerR 1 wrx2 @ wlrnat_yz_x ehlnat_1p_x2 : (ehlnat_x2 [I] urnat_x2) @ 1 =
1 @ ulnat_x2 ehrnat_p1_y0 : (ehrnat_y0 [I] ulnat_y0) @ 1 =
1 @ urnat_y0 ehrnat_p1_y1 : (ehrnat_y1 [I] ulnat_y1) @ 1 =
1 @ urnat_y1 ehrnat_p1_z0 : (ehrnat_z0 [I] ulnat_z0) @ 1 =
1 @ urnat_z0 ehrnat_p1_z1 : (ehrnat_z1 [I] ulnat_z1) @ 1 =
1 @ urnat_z1
forall (ulnat_x1 : wlx1 @ 1 = 1 @ x1)
(ehlnat_1p_x1 : (ehlnat_x1 [I] urnat_x1) @ 1 =
1 @ ulnat_x1),
(fun (q : wlx0 @ 1 = 1 @ x0)
(s : (ehlnat_x0 [I] urnat_x0) @ 1 = 1 @ q) =>
(fun (b0 : wlx1 @ 1 = 1 @ wlx2)
(x : (fun s0 : wlx1 @ 1 = 1 @ wlx2 =>
whiskerR s0 1 @ (ehrnat_z1 [-] ehlnat_x2) =
(ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^) b0) =>
(fun (b1 : wlx0 @ 1 = 1 @ wlx1)
(x3 : (fun s0 : wlx0 @ 1 = 1 @ wlx1 =>
whiskerR s0 1 @ (ehrnat_y1 [-] ehlnat_x1) =
(ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) b1) =>
(fun (q0 : wlx0 @ (1 @ 1 ) = (1 @ 1 ) @ wlx2)
(s0 : (b1 [I] b0) @ whiskerR 1 wlx2 =
whiskerL wlx0 1 @ q0) =>
P b a b a b wlx0 x0 wrx0 wlx1 x1 wrx1 wlx2 x2 wrx2
wly0 y0 1 wly1 y1 1 wlz0 z0 1 wlz1 z1 1
(wly0 @ wlz0) (1 @ 1 )
(wly1 @ wlz1) (1 @ 1 ) q urnat_x0 ulnat_x1
urnat_x1 ulnat_x2 urnat_x2 ulnat_y0 urnat_y0
ulnat_y1 urnat_y1 ulnat_z0 urnat_z0 ulnat_z1
urnat_z1 ulnat_yz0 urnat_yz0 ulnat_yz1 urnat_yz1
ehlnat_x0 ehlnat_x1 ehlnat_x2 ehrnat_y0
ehrnat_y1 ehrnat_z0 ehrnat_z1 ehrnat_yz0
ehrnat_yz1 b1 wlrnat_y_x b0 wlrnat_z_x q0
wlrnat_yz_x 1 1 1 1 H_ulnat_yz0 H_urnat_yz0
H_ulnat_yz1 H_urnat_yz1 H_ehrnat_yz0
H_ehrnat_yz1 s0 H_wlrnat_yz_x s ehlnat_1p_x1
ehlnat_1p_x2 ehrnat_p1_y0 ehrnat_p1_y1
ehrnat_p1_z0 ehrnat_p1_z1 x3 x)
((whiskerL wlx0 1 )^ @
((b1 [I] b0) @ whiskerR 1 wlx2))
(equiv_moveL_Mp
((whiskerL wlx0 1 )^ @
((b1 [I] b0) @ whiskerR 1 wlx2))
((b1 [I] b0) @ whiskerR 1 wlx2)
(whiskerL wlx0 1 ) 1 ))
(((concat_p1 (wlx0 @ 1 ))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
concat_p1 (1 @ wlx1))
(equiv_helper (ehrnat_y1 [-] ehlnat_x1)
((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^)
(((concat_p1 (wlx0 @ 1 ))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
concat_p1 (1 @ wlx1)) 1 ))
(((concat_p1 (wlx1 @ 1 ))^ @
(((ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^) @
(ehrnat_z1 [-] ehlnat_x2)^)) @
concat_p1 (1 @ wlx2))
(equiv_helper (ehrnat_z1 [-] ehlnat_x2)
((ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^)
(((concat_p1 (wlx1 @ 1 ))^ @
(((ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^) @
(ehrnat_z1 [-] ehlnat_x2)^)) @
concat_p1 (1 @ wlx2)) 1 ))
(ehlnat_x0 [I] urnat_x0)
(rlucancel 1 )
snapply equiv_path_ind_rlucancel. X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, x0, wrx0, wlx1, x1, wrx1, wlx2, x2, wrx2 : a = b wly0, y0 : b = b wly1, y1 : a = a wlz0, z0 : b = b wlz1, z1 : a = a urnat_x0 : wrx0 @ 1 = 1 @ x0 urnat_x1 : wrx1 @ 1 = 1 @ x1 ulnat_x2 : wlx2 @ 1 = 1 @ x2 urnat_x2 : wrx2 @ 1 = 1 @ x2 ulnat_y0 : wly0 @ 1 = 1 @ y0 urnat_y0 : 1 @ 1 = 1 @ y0ulnat_y1 : wly1 @ 1 = 1 @ y1 urnat_y1 : 1 @ 1 = 1 @ y1ulnat_z0 : wlz0 @ 1 = 1 @ z0 urnat_z0 : 1 @ 1 = 1 @ z0ulnat_z1 : wlz1 @ 1 = 1 @ z1 urnat_z1 : 1 @ 1 = 1 @ z1ulnat_yz0 : (wly0 @ wlz0) @ 1 = 1 @ (y0 @ z0) urnat_yz0 : (1 @ 1 ) @ 1 = 1 @ (y0 @ z0) ulnat_yz1 : (wly1 @ wlz1) @ 1 = 1 @ (y1 @ z1) urnat_yz1 : (1 @ 1 ) @ 1 = 1 @ (y1 @ z1) ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : 1 @ 1 = 1 @ wly0ehrnat_y1 : 1 @ 1 = 1 @ wly1ehrnat_z0 : 1 @ 1 = 1 @ wlz0ehrnat_z1 : 1 @ 1 = 1 @ wlz1ehrnat_yz0 : (1 @ 1 ) @ 1 = 1 @ (wly0 @ wlz0) ehrnat_yz1 : (1 @ 1 ) @ 1 = 1 @ (wly1 @ wlz1) wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0 wlrnat_yz_x : (wly1 @ wlz1) @ wrx2 =
wrx0 @ (wly0 @ wlz0) H_ulnat_yz0 : ulnat_y0 [-] ulnat_z0 =
whiskerR 1 1 @ ulnat_yz0 H_urnat_yz0 : urnat_y0 [-] urnat_z0 =
whiskerR 1 1 @ urnat_yz0 H_ulnat_yz1 : ulnat_y1 [-] ulnat_z1 =
whiskerR 1 1 @ ulnat_yz1 H_urnat_yz1 : urnat_y1 [-] urnat_z1 =
whiskerR 1 1 @ urnat_yz1 H_ehrnat_yz0 : (ehrnat_y0 [-] ehrnat_z0) @
whiskerL 1 1 =
whiskerR 1 1 @ ehrnat_yz0 H_ehrnat_yz1 : (ehrnat_y1 [-] ehrnat_z1) @
whiskerL 1 1 =
whiskerR 1 1 @ ehrnat_yz1 H_wlrnat_yz_x : (wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 1 =
whiskerR 1 wrx2 @ wlrnat_yz_x ehlnat_1p_x2 : (ehlnat_x2 [I] urnat_x2) @ 1 =
1 @ ulnat_x2 ehrnat_p1_y0 : (ehrnat_y0 [I] ulnat_y0) @ 1 =
1 @ urnat_y0 ehrnat_p1_y1 : (ehrnat_y1 [I] ulnat_y1) @ 1 =
1 @ urnat_y1 ehrnat_p1_z0 : (ehrnat_z0 [I] ulnat_z0) @ 1 =
1 @ urnat_z0 ehrnat_p1_z1 : (ehrnat_z1 [I] ulnat_z1) @ 1 =
1 @ urnat_z1
(fun (q : wlx1 @ 1 = 1 @ x1)
(s : (ehlnat_x1 [I] urnat_x1) @ 1 = 1 @ q) =>
(fun (q0 : wlx0 @ 1 = 1 @ x0)
(s0 : (ehlnat_x0 [I] urnat_x0) @ 1 = 1 @ q0) =>
(fun (b0 : wlx1 @ 1 = 1 @ wlx2)
(x : (fun s1 : wlx1 @ 1 = 1 @ wlx2 =>
whiskerR s1 1 @ (ehrnat_z1 [-] ehlnat_x2) =
(ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^) b0) =>
(fun (b1 : wlx0 @ 1 = 1 @ wlx1)
(x3 : (fun s1 : wlx0 @ 1 = 1 @ wlx1 =>
whiskerR s1 1 @ (ehrnat_y1 [-] ehlnat_x1) =
(ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) b1) =>
(fun (q1 : wlx0 @ (1 @ 1 ) = (1 @ 1 ) @ wlx2)
(s1 : (b1 [I] b0) @ whiskerR 1 wlx2 =
whiskerL wlx0 1 @ q1) =>
P b a b a b wlx0 x0 wrx0 wlx1 x1 wrx1 wlx2 x2
wrx2 wly0 y0 1 wly1 y1 1 wlz0 z0 1 wlz1 z1 1
(wly0 @ wlz0) (1 @ 1 )
(wly1 @ wlz1) (1 @ 1 ) q0 urnat_x0 q urnat_x1
ulnat_x2 urnat_x2 ulnat_y0 urnat_y0 ulnat_y1
urnat_y1 ulnat_z0 urnat_z0 ulnat_z1 urnat_z1
ulnat_yz0 urnat_yz0 ulnat_yz1 urnat_yz1
ehlnat_x0 ehlnat_x1 ehlnat_x2 ehrnat_y0
ehrnat_y1 ehrnat_z0 ehrnat_z1 ehrnat_yz0
ehrnat_yz1 b1 wlrnat_y_x b0 wlrnat_z_x q1
wlrnat_yz_x 1 1 1 1 H_ulnat_yz0 H_urnat_yz0
H_ulnat_yz1 H_urnat_yz1 H_ehrnat_yz0
H_ehrnat_yz1 s1 H_wlrnat_yz_x s0 s ehlnat_1p_x2
ehrnat_p1_y0 ehrnat_p1_y1 ehrnat_p1_z0
ehrnat_p1_z1 x3 x)
((whiskerL wlx0 1 )^ @
((b1 [I] b0) @ whiskerR 1 wlx2))
(equiv_moveL_Mp
((whiskerL wlx0 1 )^ @
((b1 [I] b0) @ whiskerR 1 wlx2))
((b1 [I] b0) @ whiskerR 1 wlx2)
(whiskerL wlx0 1 ) 1 ))
(((concat_p1 (wlx0 @ 1 ))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
concat_p1 (1 @ wlx1))
(equiv_helper (ehrnat_y1 [-] ehlnat_x1)
((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^)
(((concat_p1 (wlx0 @ 1 ))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
concat_p1 (1 @ wlx1)) 1 ))
(((concat_p1 (wlx1 @ 1 ))^ @
(((ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^) @
(ehrnat_z1 [-] ehlnat_x2)^)) @
concat_p1 (1 @ wlx2))
(equiv_helper (ehrnat_z1 [-] ehlnat_x2)
((ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^)
(((concat_p1 (wlx1 @ 1 ))^ @
(((ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^) @
(ehrnat_z1 [-] ehlnat_x2)^)) @
concat_p1 (1 @ wlx2)) 1 ))
(ehlnat_x0 [I] urnat_x0)
(rlucancel 1 )) (ehlnat_x1 [I] urnat_x1)
(rlucancel 1 )
revert ulnat_x2 ehlnat_1p_x2.X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, x0, wrx0, wlx1, x1, wrx1, wlx2, x2, wrx2 : a = b wly0, y0 : b = b wly1, y1 : a = a wlz0, z0 : b = b wlz1, z1 : a = a urnat_x0 : wrx0 @ 1 = 1 @ x0 urnat_x1 : wrx1 @ 1 = 1 @ x1 urnat_x2 : wrx2 @ 1 = 1 @ x2 ulnat_y0 : wly0 @ 1 = 1 @ y0 urnat_y0 : 1 @ 1 = 1 @ y0ulnat_y1 : wly1 @ 1 = 1 @ y1 urnat_y1 : 1 @ 1 = 1 @ y1ulnat_z0 : wlz0 @ 1 = 1 @ z0 urnat_z0 : 1 @ 1 = 1 @ z0ulnat_z1 : wlz1 @ 1 = 1 @ z1 urnat_z1 : 1 @ 1 = 1 @ z1ulnat_yz0 : (wly0 @ wlz0) @ 1 = 1 @ (y0 @ z0) urnat_yz0 : (1 @ 1 ) @ 1 = 1 @ (y0 @ z0) ulnat_yz1 : (wly1 @ wlz1) @ 1 = 1 @ (y1 @ z1) urnat_yz1 : (1 @ 1 ) @ 1 = 1 @ (y1 @ z1) ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : 1 @ 1 = 1 @ wly0ehrnat_y1 : 1 @ 1 = 1 @ wly1ehrnat_z0 : 1 @ 1 = 1 @ wlz0ehrnat_z1 : 1 @ 1 = 1 @ wlz1ehrnat_yz0 : (1 @ 1 ) @ 1 = 1 @ (wly0 @ wlz0) ehrnat_yz1 : (1 @ 1 ) @ 1 = 1 @ (wly1 @ wlz1) wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0 wlrnat_yz_x : (wly1 @ wlz1) @ wrx2 =
wrx0 @ (wly0 @ wlz0) H_ulnat_yz0 : ulnat_y0 [-] ulnat_z0 =
whiskerR 1 1 @ ulnat_yz0 H_urnat_yz0 : urnat_y0 [-] urnat_z0 =
whiskerR 1 1 @ urnat_yz0 H_ulnat_yz1 : ulnat_y1 [-] ulnat_z1 =
whiskerR 1 1 @ ulnat_yz1 H_urnat_yz1 : urnat_y1 [-] urnat_z1 =
whiskerR 1 1 @ urnat_yz1 H_ehrnat_yz0 : (ehrnat_y0 [-] ehrnat_z0) @
whiskerL 1 1 =
whiskerR 1 1 @ ehrnat_yz0 H_ehrnat_yz1 : (ehrnat_y1 [-] ehrnat_z1) @
whiskerL 1 1 =
whiskerR 1 1 @ ehrnat_yz1 H_wlrnat_yz_x : (wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 1 =
whiskerR 1 wrx2 @ wlrnat_yz_x ehrnat_p1_y0 : (ehrnat_y0 [I] ulnat_y0) @ 1 =
1 @ urnat_y0 ehrnat_p1_y1 : (ehrnat_y1 [I] ulnat_y1) @ 1 =
1 @ urnat_y1 ehrnat_p1_z0 : (ehrnat_z0 [I] ulnat_z0) @ 1 =
1 @ urnat_z0 ehrnat_p1_z1 : (ehrnat_z1 [I] ulnat_z1) @ 1 =
1 @ urnat_z1
forall (ulnat_x2 : wlx2 @ 1 = 1 @ x2)
(ehlnat_1p_x2 : (ehlnat_x2 [I] urnat_x2) @ 1 =
1 @ ulnat_x2),
(fun (q : wlx1 @ 1 = 1 @ x1)
(s : (ehlnat_x1 [I] urnat_x1) @ 1 = 1 @ q) =>
(fun (q0 : wlx0 @ 1 = 1 @ x0)
(s0 : (ehlnat_x0 [I] urnat_x0) @ 1 = 1 @ q0) =>
(fun (b0 : wlx1 @ 1 = 1 @ wlx2)
(x : (fun s1 : wlx1 @ 1 = 1 @ wlx2 =>
whiskerR s1 1 @ (ehrnat_z1 [-] ehlnat_x2) =
(ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^) b0) =>
(fun (b1 : wlx0 @ 1 = 1 @ wlx1)
(x3 : (fun s1 : wlx0 @ 1 = 1 @ wlx1 =>
whiskerR s1 1 @ (ehrnat_y1 [-] ehlnat_x1) =
(ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) b1) =>
(fun (q1 : wlx0 @ (1 @ 1 ) = (1 @ 1 ) @ wlx2)
(s1 : (b1 [I] b0) @ whiskerR 1 wlx2 =
whiskerL wlx0 1 @ q1) =>
P b a b a b wlx0 x0 wrx0 wlx1 x1 wrx1 wlx2 x2
wrx2 wly0 y0 1 wly1 y1 1 wlz0 z0 1 wlz1 z1 1
(wly0 @ wlz0) (1 @ 1 )
(wly1 @ wlz1) (1 @ 1 ) q0 urnat_x0 q urnat_x1
ulnat_x2 urnat_x2 ulnat_y0 urnat_y0 ulnat_y1
urnat_y1 ulnat_z0 urnat_z0 ulnat_z1 urnat_z1
ulnat_yz0 urnat_yz0 ulnat_yz1 urnat_yz1
ehlnat_x0 ehlnat_x1 ehlnat_x2 ehrnat_y0
ehrnat_y1 ehrnat_z0 ehrnat_z1 ehrnat_yz0
ehrnat_yz1 b1 wlrnat_y_x b0 wlrnat_z_x q1
wlrnat_yz_x 1 1 1 1 H_ulnat_yz0 H_urnat_yz0
H_ulnat_yz1 H_urnat_yz1 H_ehrnat_yz0
H_ehrnat_yz1 s1 H_wlrnat_yz_x s0 s ehlnat_1p_x2
ehrnat_p1_y0 ehrnat_p1_y1 ehrnat_p1_z0
ehrnat_p1_z1 x3 x)
((whiskerL wlx0 1 )^ @
((b1 [I] b0) @ whiskerR 1 wlx2))
(equiv_moveL_Mp
((whiskerL wlx0 1 )^ @
((b1 [I] b0) @ whiskerR 1 wlx2))
((b1 [I] b0) @ whiskerR 1 wlx2)
(whiskerL wlx0 1 ) 1 ))
(((concat_p1 (wlx0 @ 1 ))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
concat_p1 (1 @ wlx1))
(equiv_helper (ehrnat_y1 [-] ehlnat_x1)
((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^)
(((concat_p1 (wlx0 @ 1 ))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
concat_p1 (1 @ wlx1)) 1 ))
(((concat_p1 (wlx1 @ 1 ))^ @
(((ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^) @
(ehrnat_z1 [-] ehlnat_x2)^)) @
concat_p1 (1 @ wlx2))
(equiv_helper (ehrnat_z1 [-] ehlnat_x2)
((ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^)
(((concat_p1 (wlx1 @ 1 ))^ @
(((ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^) @
(ehrnat_z1 [-] ehlnat_x2)^)) @
concat_p1 (1 @ wlx2)) 1 ))
(ehlnat_x0 [I] urnat_x0)
(rlucancel 1 )) (ehlnat_x1 [I] urnat_x1)
(rlucancel 1 )
snapply equiv_path_ind_rlucancel. X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, x0, wrx0, wlx1, x1, wrx1, wlx2, x2, wrx2 : a = b wly0, y0 : b = b wly1, y1 : a = a wlz0, z0 : b = b wlz1, z1 : a = a urnat_x0 : wrx0 @ 1 = 1 @ x0 urnat_x1 : wrx1 @ 1 = 1 @ x1 urnat_x2 : wrx2 @ 1 = 1 @ x2 ulnat_y0 : wly0 @ 1 = 1 @ y0 urnat_y0 : 1 @ 1 = 1 @ y0ulnat_y1 : wly1 @ 1 = 1 @ y1 urnat_y1 : 1 @ 1 = 1 @ y1ulnat_z0 : wlz0 @ 1 = 1 @ z0 urnat_z0 : 1 @ 1 = 1 @ z0ulnat_z1 : wlz1 @ 1 = 1 @ z1 urnat_z1 : 1 @ 1 = 1 @ z1ulnat_yz0 : (wly0 @ wlz0) @ 1 = 1 @ (y0 @ z0) urnat_yz0 : (1 @ 1 ) @ 1 = 1 @ (y0 @ z0) ulnat_yz1 : (wly1 @ wlz1) @ 1 = 1 @ (y1 @ z1) urnat_yz1 : (1 @ 1 ) @ 1 = 1 @ (y1 @ z1) ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : 1 @ 1 = 1 @ wly0ehrnat_y1 : 1 @ 1 = 1 @ wly1ehrnat_z0 : 1 @ 1 = 1 @ wlz0ehrnat_z1 : 1 @ 1 = 1 @ wlz1ehrnat_yz0 : (1 @ 1 ) @ 1 = 1 @ (wly0 @ wlz0) ehrnat_yz1 : (1 @ 1 ) @ 1 = 1 @ (wly1 @ wlz1) wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0 wlrnat_yz_x : (wly1 @ wlz1) @ wrx2 =
wrx0 @ (wly0 @ wlz0) H_ulnat_yz0 : ulnat_y0 [-] ulnat_z0 =
whiskerR 1 1 @ ulnat_yz0 H_urnat_yz0 : urnat_y0 [-] urnat_z0 =
whiskerR 1 1 @ urnat_yz0 H_ulnat_yz1 : ulnat_y1 [-] ulnat_z1 =
whiskerR 1 1 @ ulnat_yz1 H_urnat_yz1 : urnat_y1 [-] urnat_z1 =
whiskerR 1 1 @ urnat_yz1 H_ehrnat_yz0 : (ehrnat_y0 [-] ehrnat_z0) @
whiskerL 1 1 =
whiskerR 1 1 @ ehrnat_yz0 H_ehrnat_yz1 : (ehrnat_y1 [-] ehrnat_z1) @
whiskerL 1 1 =
whiskerR 1 1 @ ehrnat_yz1 H_wlrnat_yz_x : (wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 1 =
whiskerR 1 wrx2 @ wlrnat_yz_x ehrnat_p1_y0 : (ehrnat_y0 [I] ulnat_y0) @ 1 =
1 @ urnat_y0 ehrnat_p1_y1 : (ehrnat_y1 [I] ulnat_y1) @ 1 =
1 @ urnat_y1 ehrnat_p1_z0 : (ehrnat_z0 [I] ulnat_z0) @ 1 =
1 @ urnat_z0 ehrnat_p1_z1 : (ehrnat_z1 [I] ulnat_z1) @ 1 =
1 @ urnat_z1
(fun (q : wlx2 @ 1 = 1 @ x2)
(s : (ehlnat_x2 [I] urnat_x2) @ 1 = 1 @ q) =>
(fun (q0 : wlx1 @ 1 = 1 @ x1)
(s0 : (ehlnat_x1 [I] urnat_x1) @ 1 = 1 @ q0) =>
(fun (q1 : wlx0 @ 1 = 1 @ x0)
(s1 : (ehlnat_x0 [I] urnat_x0) @ 1 = 1 @ q1) =>
(fun (b0 : wlx1 @ 1 = 1 @ wlx2)
(x : (fun s2 : wlx1 @ 1 = 1 @ wlx2 =>
whiskerR s2 1 @ (ehrnat_z1 [-] ehlnat_x2) =
(ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^) b0) =>
(fun (b1 : wlx0 @ 1 = 1 @ wlx1)
(x3 : (fun s2 : wlx0 @ 1 = 1 @ wlx1 =>
whiskerR s2 1 @
(ehrnat_y1 [-] ehlnat_x1) =
(ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) b1) =>
(fun (q2 : wlx0 @ (1 @ 1 ) = (1 @ 1 ) @ wlx2)
(s2 : (b1 [I] b0) @ whiskerR 1 wlx2 =
whiskerL wlx0 1 @ q2) =>
P b a b a b wlx0 x0 wrx0 wlx1 x1 wrx1 wlx2 x2
wrx2 wly0 y0 1 wly1 y1 1 wlz0 z0 1 wlz1 z1 1
(wly0 @ wlz0) (1 @ 1 )
(wly1 @ wlz1) (1 @ 1 ) q1 urnat_x0 q0 urnat_x1
q urnat_x2 ulnat_y0 urnat_y0 ulnat_y1 urnat_y1
ulnat_z0 urnat_z0 ulnat_z1 urnat_z1 ulnat_yz0
urnat_yz0 ulnat_yz1 urnat_yz1 ehlnat_x0
ehlnat_x1 ehlnat_x2 ehrnat_y0 ehrnat_y1
ehrnat_z0 ehrnat_z1 ehrnat_yz0 ehrnat_yz1 b1
wlrnat_y_x b0 wlrnat_z_x q2 wlrnat_yz_x 1 1 1
1 H_ulnat_yz0 H_urnat_yz0 H_ulnat_yz1
H_urnat_yz1 H_ehrnat_yz0 H_ehrnat_yz1 s2
H_wlrnat_yz_x s1 s0 s ehrnat_p1_y0
ehrnat_p1_y1 ehrnat_p1_z0 ehrnat_p1_z1 x3 x)
((whiskerL wlx0 1 )^ @
((b1 [I] b0) @ whiskerR 1 wlx2))
(equiv_moveL_Mp
((whiskerL wlx0 1 )^ @
((b1 [I] b0) @ whiskerR 1 wlx2))
((b1 [I] b0) @ whiskerR 1 wlx2)
(whiskerL wlx0 1 ) 1 ))
(((concat_p1 (wlx0 @ 1 ))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
concat_p1 (1 @ wlx1))
(equiv_helper (ehrnat_y1 [-] ehlnat_x1)
((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^)
(((concat_p1 (wlx0 @ 1 ))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
concat_p1 (1 @ wlx1)) 1 ))
(((concat_p1 (wlx1 @ 1 ))^ @
(((ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^) @
(ehrnat_z1 [-] ehlnat_x2)^)) @
concat_p1 (1 @ wlx2))
(equiv_helper (ehrnat_z1 [-] ehlnat_x2)
((ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^)
(((concat_p1 (wlx1 @ 1 ))^ @
(((ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^) @
(ehrnat_z1 [-] ehlnat_x2)^)) @
concat_p1 (1 @ wlx2)) 1 ))
(ehlnat_x0 [I] urnat_x0)
(rlucancel 1 )) (ehlnat_x1 [I] urnat_x1)
(rlucancel 1 )) (ehlnat_x2 [I] urnat_x2)
(rlucancel 1 )
revert urnat_yz0 H_urnat_yz0.X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, x0, wrx0, wlx1, x1, wrx1, wlx2, x2, wrx2 : a = b wly0, y0 : b = b wly1, y1 : a = a wlz0, z0 : b = b wlz1, z1 : a = a urnat_x0 : wrx0 @ 1 = 1 @ x0 urnat_x1 : wrx1 @ 1 = 1 @ x1 urnat_x2 : wrx2 @ 1 = 1 @ x2 ulnat_y0 : wly0 @ 1 = 1 @ y0 urnat_y0 : 1 @ 1 = 1 @ y0ulnat_y1 : wly1 @ 1 = 1 @ y1 urnat_y1 : 1 @ 1 = 1 @ y1ulnat_z0 : wlz0 @ 1 = 1 @ z0 urnat_z0 : 1 @ 1 = 1 @ z0ulnat_z1 : wlz1 @ 1 = 1 @ z1 urnat_z1 : 1 @ 1 = 1 @ z1ulnat_yz0 : (wly0 @ wlz0) @ 1 = 1 @ (y0 @ z0) ulnat_yz1 : (wly1 @ wlz1) @ 1 = 1 @ (y1 @ z1) urnat_yz1 : (1 @ 1 ) @ 1 = 1 @ (y1 @ z1) ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : 1 @ 1 = 1 @ wly0ehrnat_y1 : 1 @ 1 = 1 @ wly1ehrnat_z0 : 1 @ 1 = 1 @ wlz0ehrnat_z1 : 1 @ 1 = 1 @ wlz1ehrnat_yz0 : (1 @ 1 ) @ 1 = 1 @ (wly0 @ wlz0) ehrnat_yz1 : (1 @ 1 ) @ 1 = 1 @ (wly1 @ wlz1) wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0 wlrnat_yz_x : (wly1 @ wlz1) @ wrx2 =
wrx0 @ (wly0 @ wlz0) H_ulnat_yz0 : ulnat_y0 [-] ulnat_z0 =
whiskerR 1 1 @ ulnat_yz0 H_ulnat_yz1 : ulnat_y1 [-] ulnat_z1 =
whiskerR 1 1 @ ulnat_yz1 H_urnat_yz1 : urnat_y1 [-] urnat_z1 =
whiskerR 1 1 @ urnat_yz1 H_ehrnat_yz0 : (ehrnat_y0 [-] ehrnat_z0) @
whiskerL 1 1 =
whiskerR 1 1 @ ehrnat_yz0 H_ehrnat_yz1 : (ehrnat_y1 [-] ehrnat_z1) @
whiskerL 1 1 =
whiskerR 1 1 @ ehrnat_yz1 H_wlrnat_yz_x : (wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 1 =
whiskerR 1 wrx2 @ wlrnat_yz_x ehrnat_p1_y0 : (ehrnat_y0 [I] ulnat_y0) @ 1 =
1 @ urnat_y0 ehrnat_p1_y1 : (ehrnat_y1 [I] ulnat_y1) @ 1 =
1 @ urnat_y1 ehrnat_p1_z0 : (ehrnat_z0 [I] ulnat_z0) @ 1 =
1 @ urnat_z0 ehrnat_p1_z1 : (ehrnat_z1 [I] ulnat_z1) @ 1 =
1 @ urnat_z1
forall (urnat_yz0 : (1 @ 1 ) @ 1 = 1 @ (y0 @ z0))
(H_urnat_yz0 : urnat_y0 [-] urnat_z0 =
whiskerR 1 1 @ urnat_yz0),
(fun (q : wlx2 @ 1 = 1 @ x2)
(s : (ehlnat_x2 [I] urnat_x2) @ 1 = 1 @ q) =>
(fun (q0 : wlx1 @ 1 = 1 @ x1)
(s0 : (ehlnat_x1 [I] urnat_x1) @ 1 = 1 @ q0) =>
(fun (q1 : wlx0 @ 1 = 1 @ x0)
(s1 : (ehlnat_x0 [I] urnat_x0) @ 1 = 1 @ q1) =>
(fun (b0 : wlx1 @ 1 = 1 @ wlx2)
(x : (fun s2 : wlx1 @ 1 = 1 @ wlx2 =>
whiskerR s2 1 @ (ehrnat_z1 [-] ehlnat_x2) =
(ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^) b0) =>
(fun (b1 : wlx0 @ 1 = 1 @ wlx1)
(x3 : (fun s2 : wlx0 @ 1 = 1 @ wlx1 =>
whiskerR s2 1 @
(ehrnat_y1 [-] ehlnat_x1) =
(ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) b1) =>
(fun (q2 : wlx0 @ (1 @ 1 ) = (1 @ 1 ) @ wlx2)
(s2 : (b1 [I] b0) @ whiskerR 1 wlx2 =
whiskerL wlx0 1 @ q2) =>
P b a b a b wlx0 x0 wrx0 wlx1 x1 wrx1 wlx2 x2
wrx2 wly0 y0 1 wly1 y1 1 wlz0 z0 1 wlz1 z1 1
(wly0 @ wlz0) (1 @ 1 )
(wly1 @ wlz1) (1 @ 1 ) q1 urnat_x0 q0 urnat_x1
q urnat_x2 ulnat_y0 urnat_y0 ulnat_y1 urnat_y1
ulnat_z0 urnat_z0 ulnat_z1 urnat_z1 ulnat_yz0
urnat_yz0 ulnat_yz1 urnat_yz1 ehlnat_x0
ehlnat_x1 ehlnat_x2 ehrnat_y0 ehrnat_y1
ehrnat_z0 ehrnat_z1 ehrnat_yz0 ehrnat_yz1 b1
wlrnat_y_x b0 wlrnat_z_x q2 wlrnat_yz_x 1 1 1
1 H_ulnat_yz0 H_urnat_yz0 H_ulnat_yz1
H_urnat_yz1 H_ehrnat_yz0 H_ehrnat_yz1 s2
H_wlrnat_yz_x s1 s0 s ehrnat_p1_y0
ehrnat_p1_y1 ehrnat_p1_z0 ehrnat_p1_z1 x3 x)
((whiskerL wlx0 1 )^ @
((b1 [I] b0) @ whiskerR 1 wlx2))
(equiv_moveL_Mp
((whiskerL wlx0 1 )^ @
((b1 [I] b0) @ whiskerR 1 wlx2))
((b1 [I] b0) @ whiskerR 1 wlx2)
(whiskerL wlx0 1 ) 1 ))
(((concat_p1 (wlx0 @ 1 ))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
concat_p1 (1 @ wlx1))
(equiv_helper (ehrnat_y1 [-] ehlnat_x1)
((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^)
(((concat_p1 (wlx0 @ 1 ))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
concat_p1 (1 @ wlx1)) 1 ))
(((concat_p1 (wlx1 @ 1 ))^ @
(((ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^) @
(ehrnat_z1 [-] ehlnat_x2)^)) @
concat_p1 (1 @ wlx2))
(equiv_helper (ehrnat_z1 [-] ehlnat_x2)
((ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^)
(((concat_p1 (wlx1 @ 1 ))^ @
(((ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^) @
(ehrnat_z1 [-] ehlnat_x2)^)) @
concat_p1 (1 @ wlx2)) 1 ))
(ehlnat_x0 [I] urnat_x0)
(rlucancel 1 )) (ehlnat_x1 [I] urnat_x1)
(rlucancel 1 )) (ehlnat_x2 [I] urnat_x2)
(rlucancel 1 )
snapply equiv_path_ind_moveL_Mp. X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, x0, wrx0, wlx1, x1, wrx1, wlx2, x2, wrx2 : a = b wly0, y0 : b = b wly1, y1 : a = a wlz0, z0 : b = b wlz1, z1 : a = a urnat_x0 : wrx0 @ 1 = 1 @ x0 urnat_x1 : wrx1 @ 1 = 1 @ x1 urnat_x2 : wrx2 @ 1 = 1 @ x2 ulnat_y0 : wly0 @ 1 = 1 @ y0 urnat_y0 : 1 @ 1 = 1 @ y0ulnat_y1 : wly1 @ 1 = 1 @ y1 urnat_y1 : 1 @ 1 = 1 @ y1ulnat_z0 : wlz0 @ 1 = 1 @ z0 urnat_z0 : 1 @ 1 = 1 @ z0ulnat_z1 : wlz1 @ 1 = 1 @ z1 urnat_z1 : 1 @ 1 = 1 @ z1ulnat_yz0 : (wly0 @ wlz0) @ 1 = 1 @ (y0 @ z0) ulnat_yz1 : (wly1 @ wlz1) @ 1 = 1 @ (y1 @ z1) urnat_yz1 : (1 @ 1 ) @ 1 = 1 @ (y1 @ z1) ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : 1 @ 1 = 1 @ wly0ehrnat_y1 : 1 @ 1 = 1 @ wly1ehrnat_z0 : 1 @ 1 = 1 @ wlz0ehrnat_z1 : 1 @ 1 = 1 @ wlz1ehrnat_yz0 : (1 @ 1 ) @ 1 = 1 @ (wly0 @ wlz0) ehrnat_yz1 : (1 @ 1 ) @ 1 = 1 @ (wly1 @ wlz1) wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0 wlrnat_yz_x : (wly1 @ wlz1) @ wrx2 =
wrx0 @ (wly0 @ wlz0) H_ulnat_yz0 : ulnat_y0 [-] ulnat_z0 =
whiskerR 1 1 @ ulnat_yz0 H_ulnat_yz1 : ulnat_y1 [-] ulnat_z1 =
whiskerR 1 1 @ ulnat_yz1 H_urnat_yz1 : urnat_y1 [-] urnat_z1 =
whiskerR 1 1 @ urnat_yz1 H_ehrnat_yz0 : (ehrnat_y0 [-] ehrnat_z0) @
whiskerL 1 1 =
whiskerR 1 1 @ ehrnat_yz0 H_ehrnat_yz1 : (ehrnat_y1 [-] ehrnat_z1) @
whiskerL 1 1 =
whiskerR 1 1 @ ehrnat_yz1 H_wlrnat_yz_x : (wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 1 =
whiskerR 1 wrx2 @ wlrnat_yz_x ehrnat_p1_y0 : (ehrnat_y0 [I] ulnat_y0) @ 1 =
1 @ urnat_y0 ehrnat_p1_y1 : (ehrnat_y1 [I] ulnat_y1) @ 1 =
1 @ urnat_y1 ehrnat_p1_z0 : (ehrnat_z0 [I] ulnat_z0) @ 1 =
1 @ urnat_z0 ehrnat_p1_z1 : (ehrnat_z1 [I] ulnat_z1) @ 1 =
1 @ urnat_z1
(fun (q : (1 @ 1 ) @ 1 = 1 @ (y0 @ z0))
(s : urnat_y0 [-] urnat_z0 = whiskerR 1 1 @ q) =>
(fun (q0 : wlx2 @ 1 = 1 @ x2)
(s0 : (ehlnat_x2 [I] urnat_x2) @ 1 = 1 @ q0) =>
(fun (q1 : wlx1 @ 1 = 1 @ x1)
(s1 : (ehlnat_x1 [I] urnat_x1) @ 1 = 1 @ q1) =>
(fun (q2 : wlx0 @ 1 = 1 @ x0)
(s2 : (ehlnat_x0 [I] urnat_x0) @ 1 = 1 @ q2) =>
(fun (b0 : wlx1 @ 1 = 1 @ wlx2)
(x : (fun s3 : wlx1 @ 1 = 1 @ wlx2 =>
whiskerR s3 1 @ (ehrnat_z1 [-] ehlnat_x2) =
(ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^) b0) =>
(fun (b1 : wlx0 @ 1 = 1 @ wlx1)
(x3 : (fun s3 : wlx0 @ 1 = 1 @ wlx1 =>
whiskerR s3 1 @
(ehrnat_y1 [-] ehlnat_x1) =
(ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) b1) =>
(fun (q3 : wlx0 @ (1 @ 1 ) = (1 @ 1 ) @ wlx2)
(s3 : (b1 [I] b0) @ whiskerR 1 wlx2 =
whiskerL wlx0 1 @ q3) =>
P b a b a b wlx0 x0 wrx0 wlx1 x1 wrx1 wlx2 x2
wrx2 wly0 y0 1 wly1 y1 1 wlz0 z0 1 wlz1 z1 1
(wly0 @ wlz0) (1 @ 1 )
(wly1 @ wlz1) (1 @ 1 ) q2 urnat_x0 q1 urnat_x1
q0 urnat_x2 ulnat_y0 urnat_y0 ulnat_y1
urnat_y1 ulnat_z0 urnat_z0 ulnat_z1 urnat_z1
ulnat_yz0 q ulnat_yz1 urnat_yz1 ehlnat_x0
ehlnat_x1 ehlnat_x2 ehrnat_y0 ehrnat_y1
ehrnat_z0 ehrnat_z1 ehrnat_yz0 ehrnat_yz1 b1
wlrnat_y_x b0 wlrnat_z_x q3 wlrnat_yz_x 1 1 1
1 H_ulnat_yz0 s H_ulnat_yz1 H_urnat_yz1
H_ehrnat_yz0 H_ehrnat_yz1 s3 H_wlrnat_yz_x s2
s1 s0 ehrnat_p1_y0 ehrnat_p1_y1 ehrnat_p1_z0
ehrnat_p1_z1 x3 x)
((whiskerL wlx0 1 )^ @
((b1 [I] b0) @ whiskerR 1 wlx2))
(equiv_moveL_Mp
((whiskerL wlx0 1 )^ @
((b1 [I] b0) @ whiskerR 1 wlx2))
((b1 [I] b0) @ whiskerR 1 wlx2)
(whiskerL wlx0 1 ) 1 ))
(((concat_p1 (wlx0 @ 1 ))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
concat_p1 (1 @ wlx1))
(equiv_helper (ehrnat_y1 [-] ehlnat_x1)
((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^)
(((concat_p1 (wlx0 @ 1 ))^ @
(((...) @ whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
concat_p1 (1 @ wlx1)) 1 ))
(((concat_p1 (wlx1 @ 1 ))^ @
(((ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^) @
(ehrnat_z1 [-] ehlnat_x2)^)) @
concat_p1 (1 @ wlx2))
(equiv_helper (ehrnat_z1 [-] ehlnat_x2)
((ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^)
(((concat_p1 (wlx1 @ 1 ))^ @
(((ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^) @
(ehrnat_z1 [-] ehlnat_x2)^)) @
concat_p1 (1 @ wlx2)) 1 ))
(ehlnat_x0 [I] urnat_x0)
(rlucancel 1 )) (ehlnat_x1 [I] urnat_x1)
(rlucancel 1 )) (ehlnat_x2 [I] urnat_x2)
(rlucancel 1 ))
((whiskerR 1 1 )^ @ (urnat_y0 [-] urnat_z0))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @ (urnat_y0 [-] urnat_z0))
(urnat_y0 [-] urnat_z0)
(whiskerR 1 1 ) 1 )
revert urnat_yz1 H_urnat_yz1.X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, x0, wrx0, wlx1, x1, wrx1, wlx2, x2, wrx2 : a = b wly0, y0 : b = b wly1, y1 : a = a wlz0, z0 : b = b wlz1, z1 : a = a urnat_x0 : wrx0 @ 1 = 1 @ x0 urnat_x1 : wrx1 @ 1 = 1 @ x1 urnat_x2 : wrx2 @ 1 = 1 @ x2 ulnat_y0 : wly0 @ 1 = 1 @ y0 urnat_y0 : 1 @ 1 = 1 @ y0ulnat_y1 : wly1 @ 1 = 1 @ y1 urnat_y1 : 1 @ 1 = 1 @ y1ulnat_z0 : wlz0 @ 1 = 1 @ z0 urnat_z0 : 1 @ 1 = 1 @ z0ulnat_z1 : wlz1 @ 1 = 1 @ z1 urnat_z1 : 1 @ 1 = 1 @ z1ulnat_yz0 : (wly0 @ wlz0) @ 1 = 1 @ (y0 @ z0) ulnat_yz1 : (wly1 @ wlz1) @ 1 = 1 @ (y1 @ z1) ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : 1 @ 1 = 1 @ wly0ehrnat_y1 : 1 @ 1 = 1 @ wly1ehrnat_z0 : 1 @ 1 = 1 @ wlz0ehrnat_z1 : 1 @ 1 = 1 @ wlz1ehrnat_yz0 : (1 @ 1 ) @ 1 = 1 @ (wly0 @ wlz0) ehrnat_yz1 : (1 @ 1 ) @ 1 = 1 @ (wly1 @ wlz1) wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0 wlrnat_yz_x : (wly1 @ wlz1) @ wrx2 =
wrx0 @ (wly0 @ wlz0) H_ulnat_yz0 : ulnat_y0 [-] ulnat_z0 =
whiskerR 1 1 @ ulnat_yz0 H_ulnat_yz1 : ulnat_y1 [-] ulnat_z1 =
whiskerR 1 1 @ ulnat_yz1 H_ehrnat_yz0 : (ehrnat_y0 [-] ehrnat_z0) @
whiskerL 1 1 =
whiskerR 1 1 @ ehrnat_yz0 H_ehrnat_yz1 : (ehrnat_y1 [-] ehrnat_z1) @
whiskerL 1 1 =
whiskerR 1 1 @ ehrnat_yz1 H_wlrnat_yz_x : (wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 1 =
whiskerR 1 wrx2 @ wlrnat_yz_x ehrnat_p1_y0 : (ehrnat_y0 [I] ulnat_y0) @ 1 =
1 @ urnat_y0 ehrnat_p1_y1 : (ehrnat_y1 [I] ulnat_y1) @ 1 =
1 @ urnat_y1 ehrnat_p1_z0 : (ehrnat_z0 [I] ulnat_z0) @ 1 =
1 @ urnat_z0 ehrnat_p1_z1 : (ehrnat_z1 [I] ulnat_z1) @ 1 =
1 @ urnat_z1
forall (urnat_yz1 : (1 @ 1 ) @ 1 = 1 @ (y1 @ z1))
(H_urnat_yz1 : urnat_y1 [-] urnat_z1 =
whiskerR 1 1 @ urnat_yz1),
(fun (q : (1 @ 1 ) @ 1 = 1 @ (y0 @ z0))
(s : urnat_y0 [-] urnat_z0 = whiskerR 1 1 @ q) =>
(fun (q0 : wlx2 @ 1 = 1 @ x2)
(s0 : (ehlnat_x2 [I] urnat_x2) @ 1 = 1 @ q0) =>
(fun (q1 : wlx1 @ 1 = 1 @ x1)
(s1 : (ehlnat_x1 [I] urnat_x1) @ 1 = 1 @ q1) =>
(fun (q2 : wlx0 @ 1 = 1 @ x0)
(s2 : (ehlnat_x0 [I] urnat_x0) @ 1 = 1 @ q2) =>
(fun (b0 : wlx1 @ 1 = 1 @ wlx2)
(x : (fun s3 : wlx1 @ 1 = 1 @ wlx2 =>
whiskerR s3 1 @ (ehrnat_z1 [-] ehlnat_x2) =
(ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^) b0) =>
(fun (b1 : wlx0 @ 1 = 1 @ wlx1)
(x3 : (fun s3 : ... = ... =>
whiskerR s3 1 @ (...) =
(...) @ whiskerL 1 wlrnat_y_x^) b1) =>
(fun (q3 : wlx0 @ (...) = (...) @ wlx2)
(s3 : (...) @ whiskerR 1 wlx2 =
whiskerL wlx0 1 @ q3) =>
P b a b a b wlx0 x0 wrx0 wlx1 x1 wrx1 wlx2 x2
wrx2 wly0 y0 1 wly1 y1 1 wlz0 z0 1 wlz1 z1 1
(wly0 @ wlz0) (1 @ 1 )
(wly1 @ wlz1) (1 @ 1 ) q2 urnat_x0 q1 urnat_x1
q0 urnat_x2 ulnat_y0 urnat_y0 ulnat_y1
urnat_y1 ulnat_z0 urnat_z0 ulnat_z1 urnat_z1
ulnat_yz0 q ulnat_yz1 urnat_yz1 ehlnat_x0
ehlnat_x1 ehlnat_x2 ehrnat_y0 ehrnat_y1
ehrnat_z0 ehrnat_z1 ehrnat_yz0 ehrnat_yz1 b1
wlrnat_y_x b0 wlrnat_z_x q3 wlrnat_yz_x 1 1 1
1 H_ulnat_yz0 s H_ulnat_yz1 H_urnat_yz1
H_ehrnat_yz0 H_ehrnat_yz1 s3 H_wlrnat_yz_x s2
s1 s0 ehrnat_p1_y0 ehrnat_p1_y1 ehrnat_p1_z0
ehrnat_p1_z1 x3 x)
((whiskerL wlx0 1 )^ @
((b1 [I] b0) @ whiskerR 1 wlx2))
(equiv_moveL_Mp
((whiskerL wlx0 1 )^ @
((...) @ whiskerR 1 wlx2))
((b1 [I] b0) @ whiskerR 1 wlx2)
(whiskerL wlx0 1 ) 1 ))
(((concat_p1 (wlx0 @ 1 ))^ @
(((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^) @
(ehrnat_y1 [-] ehlnat_x1)^)) @
concat_p1 (1 @ wlx1))
(equiv_helper (ehrnat_y1 [-] ehlnat_x1)
((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^)
(((concat_p1 (wlx0 @ 1 ))^ @
((... @ ...) @ (ehrnat_y1 [-] ehlnat_x1)^)) @
concat_p1 (1 @ wlx1)) 1 ))
(((concat_p1 (wlx1 @ 1 ))^ @
(((ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^) @
(ehrnat_z1 [-] ehlnat_x2)^)) @
concat_p1 (1 @ wlx2))
(equiv_helper (ehrnat_z1 [-] ehlnat_x2)
((ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^)
(((concat_p1 (wlx1 @ 1 ))^ @
(((ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^) @
(ehrnat_z1 [-] ehlnat_x2)^)) @
concat_p1 (1 @ wlx2)) 1 ))
(ehlnat_x0 [I] urnat_x0)
(rlucancel 1 )) (ehlnat_x1 [I] urnat_x1)
(rlucancel 1 )) (ehlnat_x2 [I] urnat_x2)
(rlucancel 1 ))
((whiskerR 1 1 )^ @ (urnat_y0 [-] urnat_z0))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @ (urnat_y0 [-] urnat_z0))
(urnat_y0 [-] urnat_z0)
(whiskerR 1 1 ) 1 )
snapply equiv_path_ind_moveL_Mp. X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, x0, wrx0, wlx1, x1, wrx1, wlx2, x2, wrx2 : a = b wly0, y0 : b = b wly1, y1 : a = a wlz0, z0 : b = b wlz1, z1 : a = a urnat_x0 : wrx0 @ 1 = 1 @ x0 urnat_x1 : wrx1 @ 1 = 1 @ x1 urnat_x2 : wrx2 @ 1 = 1 @ x2 ulnat_y0 : wly0 @ 1 = 1 @ y0 urnat_y0 : 1 @ 1 = 1 @ y0ulnat_y1 : wly1 @ 1 = 1 @ y1 urnat_y1 : 1 @ 1 = 1 @ y1ulnat_z0 : wlz0 @ 1 = 1 @ z0 urnat_z0 : 1 @ 1 = 1 @ z0ulnat_z1 : wlz1 @ 1 = 1 @ z1 urnat_z1 : 1 @ 1 = 1 @ z1ulnat_yz0 : (wly0 @ wlz0) @ 1 = 1 @ (y0 @ z0) ulnat_yz1 : (wly1 @ wlz1) @ 1 = 1 @ (y1 @ z1) ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : 1 @ 1 = 1 @ wly0ehrnat_y1 : 1 @ 1 = 1 @ wly1ehrnat_z0 : 1 @ 1 = 1 @ wlz0ehrnat_z1 : 1 @ 1 = 1 @ wlz1ehrnat_yz0 : (1 @ 1 ) @ 1 = 1 @ (wly0 @ wlz0) ehrnat_yz1 : (1 @ 1 ) @ 1 = 1 @ (wly1 @ wlz1) wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0 wlrnat_yz_x : (wly1 @ wlz1) @ wrx2 =
wrx0 @ (wly0 @ wlz0) H_ulnat_yz0 : ulnat_y0 [-] ulnat_z0 =
whiskerR 1 1 @ ulnat_yz0 H_ulnat_yz1 : ulnat_y1 [-] ulnat_z1 =
whiskerR 1 1 @ ulnat_yz1 H_ehrnat_yz0 : (ehrnat_y0 [-] ehrnat_z0) @
whiskerL 1 1 =
whiskerR 1 1 @ ehrnat_yz0 H_ehrnat_yz1 : (ehrnat_y1 [-] ehrnat_z1) @
whiskerL 1 1 =
whiskerR 1 1 @ ehrnat_yz1 H_wlrnat_yz_x : (wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 1 =
whiskerR 1 wrx2 @ wlrnat_yz_x ehrnat_p1_y0 : (ehrnat_y0 [I] ulnat_y0) @ 1 =
1 @ urnat_y0 ehrnat_p1_y1 : (ehrnat_y1 [I] ulnat_y1) @ 1 =
1 @ urnat_y1 ehrnat_p1_z0 : (ehrnat_z0 [I] ulnat_z0) @ 1 =
1 @ urnat_z0 ehrnat_p1_z1 : (ehrnat_z1 [I] ulnat_z1) @ 1 =
1 @ urnat_z1
(fun (q : (1 @ 1 ) @ 1 = 1 @ (y1 @ z1))
(s : urnat_y1 [-] urnat_z1 = whiskerR 1 1 @ q) =>
(fun (q0 : (1 @ 1 ) @ 1 = 1 @ (y0 @ z0))
(s0 : urnat_y0 [-] urnat_z0 = whiskerR 1 1 @ q0)
=>
(fun (q1 : wlx2 @ 1 = 1 @ x2)
(s1 : (ehlnat_x2 [I] urnat_x2) @ 1 = 1 @ q1) =>
(fun (q2 : wlx1 @ 1 = 1 @ x1)
(s2 : (ehlnat_x1 [I] urnat_x1) @ 1 = 1 @ q2) =>
(fun (q3 : wlx0 @ 1 = 1 @ x0)
(s3 : (ehlnat_x0 [I] urnat_x0) @ 1 = 1 @ q3) =>
(fun (b0 : wlx1 @ 1 = 1 @ wlx2)
(x : (fun s4 : wlx1 @ 1 = 1 @ wlx2 =>
whiskerR s4 1 @
(ehrnat_z1 [-] ehlnat_x2) =
(ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^) b0) =>
(fun (b1 : wlx0 @ 1 = 1 @ wlx1)
(x3 : (fun ... => ... = ...) b1) =>
(fun (q4 : ... = ...) (s4 : ... = ...) =>
P b a b a b wlx0 x0 wrx0 wlx1 x1 wrx1 wlx2 x2
wrx2 wly0 y0 1 wly1 y1 1 wlz0 z0 1 wlz1 z1 1
(wly0 @ wlz0) (1 @ 1 )
(wly1 @ wlz1) (1 @ 1 ) q3 urnat_x0 q2
urnat_x1 q1 urnat_x2 ulnat_y0 urnat_y0
ulnat_y1 urnat_y1 ulnat_z0 urnat_z0 ulnat_z1
urnat_z1 ulnat_yz0 q0 ulnat_yz1 q ehlnat_x0
ehlnat_x1 ehlnat_x2 ehrnat_y0 ehrnat_y1
ehrnat_z0 ehrnat_z1 ehrnat_yz0 ehrnat_yz1 b1
wlrnat_y_x b0 wlrnat_z_x q4 wlrnat_yz_x 1 1
1 1 H_ulnat_yz0 s0 H_ulnat_yz1 s
H_ehrnat_yz0 H_ehrnat_yz1 s4 H_wlrnat_yz_x
s3 s2 s1 ehrnat_p1_y0 ehrnat_p1_y1
ehrnat_p1_z0 ehrnat_p1_z1 x3 x)
((whiskerL wlx0 1 )^ @
((...) @ whiskerR 1 wlx2))
(equiv_moveL_Mp
((...)^ @ (...))
((...) @ whiskerR 1 wlx2)
(whiskerL wlx0 1 ) 1 ))
(((concat_p1 (wlx0 @ 1 ))^ @
((... @ ...) @ (ehrnat_y1 [-] ehlnat_x1)^)) @
concat_p1 (1 @ wlx1))
(equiv_helper (ehrnat_y1 [-] ehlnat_x1)
((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^)
(((concat_p1 ...)^ @ (... @ ...^)) @
concat_p1 (1 @ wlx1)) 1 ))
(((concat_p1 (wlx1 @ 1 ))^ @
(((ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^) @
(ehrnat_z1 [-] ehlnat_x2)^)) @
concat_p1 (1 @ wlx2))
(equiv_helper (ehrnat_z1 [-] ehlnat_x2)
((ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^)
(((concat_p1 (wlx1 @ 1 ))^ @
(((...) @ whiskerL 1 wlrnat_z_x^) @
(ehrnat_z1 [-] ehlnat_x2)^)) @
concat_p1 (1 @ wlx2)) 1 ))
(ehlnat_x0 [I] urnat_x0)
(rlucancel 1 )) (ehlnat_x1 [I] urnat_x1)
(rlucancel 1 )) (ehlnat_x2 [I] urnat_x2)
(rlucancel 1 ))
((whiskerR 1 1 )^ @ (urnat_y0 [-] urnat_z0))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @ (urnat_y0 [-] urnat_z0))
(urnat_y0 [-] urnat_z0)
(whiskerR 1 1 ) 1 ))
((whiskerR 1 1 )^ @ (urnat_y1 [-] urnat_z1))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @ (urnat_y1 [-] urnat_z1))
(urnat_y1 [-] urnat_z1)
(whiskerR 1 1 ) 1 )
revert wlrnat_yz_x H_wlrnat_yz_x.X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, x0, wrx0, wlx1, x1, wrx1, wlx2, x2, wrx2 : a = b wly0, y0 : b = b wly1, y1 : a = a wlz0, z0 : b = b wlz1, z1 : a = a urnat_x0 : wrx0 @ 1 = 1 @ x0 urnat_x1 : wrx1 @ 1 = 1 @ x1 urnat_x2 : wrx2 @ 1 = 1 @ x2 ulnat_y0 : wly0 @ 1 = 1 @ y0 urnat_y0 : 1 @ 1 = 1 @ y0ulnat_y1 : wly1 @ 1 = 1 @ y1 urnat_y1 : 1 @ 1 = 1 @ y1ulnat_z0 : wlz0 @ 1 = 1 @ z0 urnat_z0 : 1 @ 1 = 1 @ z0ulnat_z1 : wlz1 @ 1 = 1 @ z1 urnat_z1 : 1 @ 1 = 1 @ z1ulnat_yz0 : (wly0 @ wlz0) @ 1 = 1 @ (y0 @ z0) ulnat_yz1 : (wly1 @ wlz1) @ 1 = 1 @ (y1 @ z1) ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : 1 @ 1 = 1 @ wly0ehrnat_y1 : 1 @ 1 = 1 @ wly1ehrnat_z0 : 1 @ 1 = 1 @ wlz0ehrnat_z1 : 1 @ 1 = 1 @ wlz1ehrnat_yz0 : (1 @ 1 ) @ 1 = 1 @ (wly0 @ wlz0) ehrnat_yz1 : (1 @ 1 ) @ 1 = 1 @ (wly1 @ wlz1) wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0 H_ulnat_yz0 : ulnat_y0 [-] ulnat_z0 =
whiskerR 1 1 @ ulnat_yz0 H_ulnat_yz1 : ulnat_y1 [-] ulnat_z1 =
whiskerR 1 1 @ ulnat_yz1 H_ehrnat_yz0 : (ehrnat_y0 [-] ehrnat_z0) @
whiskerL 1 1 =
whiskerR 1 1 @ ehrnat_yz0 H_ehrnat_yz1 : (ehrnat_y1 [-] ehrnat_z1) @
whiskerL 1 1 =
whiskerR 1 1 @ ehrnat_yz1 ehrnat_p1_y0 : (ehrnat_y0 [I] ulnat_y0) @ 1 =
1 @ urnat_y0 ehrnat_p1_y1 : (ehrnat_y1 [I] ulnat_y1) @ 1 =
1 @ urnat_y1 ehrnat_p1_z0 : (ehrnat_z0 [I] ulnat_z0) @ 1 =
1 @ urnat_z0 ehrnat_p1_z1 : (ehrnat_z1 [I] ulnat_z1) @ 1 =
1 @ urnat_z1
forall
(wlrnat_yz_x : (wly1 @ wlz1) @ wrx2 =
wrx0 @ (wly0 @ wlz0))
(H_wlrnat_yz_x : (wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 1 =
whiskerR 1 wrx2 @ wlrnat_yz_x),
(fun (q : (1 @ 1 ) @ 1 = 1 @ (y1 @ z1))
(s : urnat_y1 [-] urnat_z1 = whiskerR 1 1 @ q) =>
(fun (q0 : (1 @ 1 ) @ 1 = 1 @ (y0 @ z0))
(s0 : urnat_y0 [-] urnat_z0 = whiskerR 1 1 @ q0)
=>
(fun (q1 : wlx2 @ 1 = 1 @ x2)
(s1 : (ehlnat_x2 [I] urnat_x2) @ 1 = 1 @ q1) =>
(fun (q2 : wlx1 @ 1 = 1 @ x1)
(s2 : (ehlnat_x1 [I] urnat_x1) @ 1 = 1 @ q2) =>
(fun (q3 : wlx0 @ 1 = 1 @ x0)
(s3 : (ehlnat_x0 [I] urnat_x0) @ 1 = 1 @ q3) =>
(fun (b0 : wlx1 @ 1 = 1 @ wlx2)
(x : (fun s4 : ... = ... =>
whiskerR s4 1 @ (...) =
(...) @ whiskerL 1 wlrnat_z_x^) b0) =>
(fun (b1 : wlx0 @ 1 = 1 @ wlx1)
(x3 : (... => ...) b1) =>
(fun (q4 : ...) (s4 : ...) =>
P b a b a b wlx0 x0 wrx0 wlx1 x1 wrx1 wlx2 x2
wrx2 wly0 y0 1 wly1 y1 1 wlz0 z0 1 wlz1 z1 1
(...) (...) (...)
(...) q3 urnat_x0 q2 urnat_x1 q1 urnat_x2
ulnat_y0 urnat_y0 ulnat_y1 urnat_y1 ulnat_z0
urnat_z0 ulnat_z1 urnat_z1 ulnat_yz0 q0
ulnat_yz1 q ehlnat_x0 ehlnat_x1 ehlnat_x2
ehrnat_y0 ehrnat_y1 ehrnat_z0 ehrnat_z1
ehrnat_yz0 ehrnat_yz1 b1 wlrnat_y_x b0
wlrnat_z_x q4 wlrnat_yz_x 1 1 1 1
H_ulnat_yz0 s0 H_ulnat_yz1 s H_ehrnat_yz0
H_ehrnat_yz1 s4 H_wlrnat_yz_x s3 s2 s1
ehrnat_p1_y0 ehrnat_p1_y1 ehrnat_p1_z0
ehrnat_p1_z1 x3 x)
((whiskerL wlx0 1 )^ @ (... @ ...))
(equiv_moveL_Mp
(...^ @ ...)
(... @ ...) (whiskerL wlx0 1 ) 1 ))
(((concat_p1 (...))^ @ ((...) @ (...)^)) @
concat_p1 (1 @ wlx1))
(equiv_helper (ehrnat_y1 [-] ehlnat_x1)
((ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^)
(((...)^ @ (...)) @ concat_p1 (1 @ wlx1)) 1 ))
(((concat_p1 (wlx1 @ 1 ))^ @
(((ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^) @
(ehrnat_z1 [-] ehlnat_x2)^)) @
concat_p1 (1 @ wlx2))
(equiv_helper (ehrnat_z1 [-] ehlnat_x2)
((ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^)
(((concat_p1 (wlx1 @ 1 ))^ @
((... @ ...) @ (ehrnat_z1 [-] ehlnat_x2)^)) @
concat_p1 (1 @ wlx2)) 1 ))
(ehlnat_x0 [I] urnat_x0)
(rlucancel 1 )) (ehlnat_x1 [I] urnat_x1)
(rlucancel 1 )) (ehlnat_x2 [I] urnat_x2)
(rlucancel 1 ))
((whiskerR 1 1 )^ @ (urnat_y0 [-] urnat_z0))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @ (urnat_y0 [-] urnat_z0))
(urnat_y0 [-] urnat_z0)
(whiskerR 1 1 ) 1 ))
((whiskerR 1 1 )^ @ (urnat_y1 [-] urnat_z1))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @ (urnat_y1 [-] urnat_z1))
(urnat_y1 [-] urnat_z1)
(whiskerR 1 1 ) 1 )
snapply equiv_path_ind_moveL_Mp. X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, x0, wrx0, wlx1, x1, wrx1, wlx2, x2, wrx2 : a = b wly0, y0 : b = b wly1, y1 : a = a wlz0, z0 : b = b wlz1, z1 : a = a urnat_x0 : wrx0 @ 1 = 1 @ x0 urnat_x1 : wrx1 @ 1 = 1 @ x1 urnat_x2 : wrx2 @ 1 = 1 @ x2 ulnat_y0 : wly0 @ 1 = 1 @ y0 urnat_y0 : 1 @ 1 = 1 @ y0ulnat_y1 : wly1 @ 1 = 1 @ y1 urnat_y1 : 1 @ 1 = 1 @ y1ulnat_z0 : wlz0 @ 1 = 1 @ z0 urnat_z0 : 1 @ 1 = 1 @ z0ulnat_z1 : wlz1 @ 1 = 1 @ z1 urnat_z1 : 1 @ 1 = 1 @ z1ulnat_yz0 : (wly0 @ wlz0) @ 1 = 1 @ (y0 @ z0) ulnat_yz1 : (wly1 @ wlz1) @ 1 = 1 @ (y1 @ z1) ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : 1 @ 1 = 1 @ wly0ehrnat_y1 : 1 @ 1 = 1 @ wly1ehrnat_z0 : 1 @ 1 = 1 @ wlz0ehrnat_z1 : 1 @ 1 = 1 @ wlz1ehrnat_yz0 : (1 @ 1 ) @ 1 = 1 @ (wly0 @ wlz0) ehrnat_yz1 : (1 @ 1 ) @ 1 = 1 @ (wly1 @ wlz1) wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0 H_ulnat_yz0 : ulnat_y0 [-] ulnat_z0 =
whiskerR 1 1 @ ulnat_yz0 H_ulnat_yz1 : ulnat_y1 [-] ulnat_z1 =
whiskerR 1 1 @ ulnat_yz1 H_ehrnat_yz0 : (ehrnat_y0 [-] ehrnat_z0) @
whiskerL 1 1 =
whiskerR 1 1 @ ehrnat_yz0 H_ehrnat_yz1 : (ehrnat_y1 [-] ehrnat_z1) @
whiskerL 1 1 =
whiskerR 1 1 @ ehrnat_yz1 ehrnat_p1_y0 : (ehrnat_y0 [I] ulnat_y0) @ 1 =
1 @ urnat_y0 ehrnat_p1_y1 : (ehrnat_y1 [I] ulnat_y1) @ 1 =
1 @ urnat_y1 ehrnat_p1_z0 : (ehrnat_z0 [I] ulnat_z0) @ 1 =
1 @ urnat_z0 ehrnat_p1_z1 : (ehrnat_z1 [I] ulnat_z1) @ 1 =
1 @ urnat_z1
(fun (q : (wly1 @ wlz1) @ wrx2 = wrx0 @ (wly0 @ wlz0))
(s : (wlrnat_y_x [-] wlrnat_z_x) @ whiskerL wrx0 1 =
whiskerR 1 wrx2 @ q) =>
(fun (q0 : (1 @ 1 ) @ 1 = 1 @ (y1 @ z1))
(s0 : urnat_y1 [-] urnat_z1 = whiskerR 1 1 @ q0)
=>
(fun (q1 : (1 @ 1 ) @ 1 = 1 @ (y0 @ z0))
(s1 : urnat_y0 [-] urnat_z0 = whiskerR 1 1 @ q1)
=>
(fun (q2 : wlx2 @ 1 = 1 @ x2)
(s2 : (ehlnat_x2 [I] urnat_x2) @ 1 = 1 @ q2) =>
(fun (q3 : wlx1 @ 1 = 1 @ x1)
(s3 : (ehlnat_x1 [I] urnat_x1) @ 1 = 1 @ q3) =>
(fun (q4 : wlx0 @ 1 = 1 @ x0)
(s4 : (ehlnat_x0 [I] urnat_x0) @ 1 = 1 @ q4)
=>
(fun (b0 : wlx1 @ 1 = 1 @ wlx2)
(x : (fun ... => ... = ...) b0) =>
(fun (b1 : ... = ...) (x3 : ... b1) =>
(... => ...) (...^ @ ...)
(equiv_moveL_Mp ... ... ... 1 ))
(((...)^ @ (...)) @ concat_p1 (1 @ wlx1))
(equiv_helper (ehrnat_y1 [-] ehlnat_x1)
((...) @ whiskerL 1 wlrnat_y_x^)
((...) @ concat_p1 ...) 1 ))
(((concat_p1 (wlx1 @ 1 ))^ @
((... @ ...) @ (ehrnat_z1 [-] ehlnat_x2)^)) @
concat_p1 (1 @ wlx2))
(equiv_helper (ehrnat_z1 [-] ehlnat_x2)
((ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^)
(((concat_p1 ...)^ @ (... @ ...^)) @
concat_p1 (1 @ wlx2)) 1 ))
(ehlnat_x0 [I] urnat_x0)
(rlucancel 1 )) (ehlnat_x1 [I] urnat_x1)
(rlucancel 1 )) (ehlnat_x2 [I] urnat_x2)
(rlucancel 1 ))
((whiskerR 1 1 )^ @ (urnat_y0 [-] urnat_z0))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @ (urnat_y0 [-] urnat_z0))
(urnat_y0 [-] urnat_z0)
(whiskerR 1 1 ) 1 ))
((whiskerR 1 1 )^ @ (urnat_y1 [-] urnat_z1))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @ (urnat_y1 [-] urnat_z1))
(urnat_y1 [-] urnat_z1)
(whiskerR 1 1 ) 1 ))
((whiskerR 1 wrx2)^ @
((wlrnat_y_x [-] wlrnat_z_x) @ whiskerL wrx0 1 ))
(equiv_moveL_Mp
((whiskerR 1 wrx2)^ @
((wlrnat_y_x [-] wlrnat_z_x) @ whiskerL wrx0 1 ))
((wlrnat_y_x [-] wlrnat_z_x) @ whiskerL wrx0 1 )
(whiskerR 1 wrx2) 1 )
revert ehrnat_yz0 H_ehrnat_yz0.X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, x0, wrx0, wlx1, x1, wrx1, wlx2, x2, wrx2 : a = b wly0, y0 : b = b wly1, y1 : a = a wlz0, z0 : b = b wlz1, z1 : a = a urnat_x0 : wrx0 @ 1 = 1 @ x0 urnat_x1 : wrx1 @ 1 = 1 @ x1 urnat_x2 : wrx2 @ 1 = 1 @ x2 ulnat_y0 : wly0 @ 1 = 1 @ y0 urnat_y0 : 1 @ 1 = 1 @ y0ulnat_y1 : wly1 @ 1 = 1 @ y1 urnat_y1 : 1 @ 1 = 1 @ y1ulnat_z0 : wlz0 @ 1 = 1 @ z0 urnat_z0 : 1 @ 1 = 1 @ z0ulnat_z1 : wlz1 @ 1 = 1 @ z1 urnat_z1 : 1 @ 1 = 1 @ z1ulnat_yz0 : (wly0 @ wlz0) @ 1 = 1 @ (y0 @ z0) ulnat_yz1 : (wly1 @ wlz1) @ 1 = 1 @ (y1 @ z1) ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : 1 @ 1 = 1 @ wly0ehrnat_y1 : 1 @ 1 = 1 @ wly1ehrnat_z0 : 1 @ 1 = 1 @ wlz0ehrnat_z1 : 1 @ 1 = 1 @ wlz1ehrnat_yz1 : (1 @ 1 ) @ 1 = 1 @ (wly1 @ wlz1) wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0 H_ulnat_yz0 : ulnat_y0 [-] ulnat_z0 =
whiskerR 1 1 @ ulnat_yz0 H_ulnat_yz1 : ulnat_y1 [-] ulnat_z1 =
whiskerR 1 1 @ ulnat_yz1 H_ehrnat_yz1 : (ehrnat_y1 [-] ehrnat_z1) @
whiskerL 1 1 =
whiskerR 1 1 @ ehrnat_yz1 ehrnat_p1_y0 : (ehrnat_y0 [I] ulnat_y0) @ 1 =
1 @ urnat_y0 ehrnat_p1_y1 : (ehrnat_y1 [I] ulnat_y1) @ 1 =
1 @ urnat_y1 ehrnat_p1_z0 : (ehrnat_z0 [I] ulnat_z0) @ 1 =
1 @ urnat_z0 ehrnat_p1_z1 : (ehrnat_z1 [I] ulnat_z1) @ 1 =
1 @ urnat_z1
forall (ehrnat_yz0 : (1 @ 1 ) @ 1 = 1 @ (wly0 @ wlz0))
(H_ehrnat_yz0 : (ehrnat_y0 [-] ehrnat_z0) @
whiskerL 1 1 =
whiskerR 1 1 @ ehrnat_yz0),
(fun (q : (wly1 @ wlz1) @ wrx2 = wrx0 @ (wly0 @ wlz0))
(s : (wlrnat_y_x [-] wlrnat_z_x) @ whiskerL wrx0 1 =
whiskerR 1 wrx2 @ q) =>
(fun (q0 : (1 @ 1 ) @ 1 = 1 @ (y1 @ z1))
(s0 : urnat_y1 [-] urnat_z1 = whiskerR 1 1 @ q0)
=>
(fun (q1 : (1 @ 1 ) @ 1 = 1 @ (y0 @ z0))
(s1 : urnat_y0 [-] urnat_z0 = whiskerR 1 1 @ q1)
=>
(fun (q2 : wlx2 @ 1 = 1 @ x2)
(s2 : (ehlnat_x2 [I] urnat_x2) @ 1 = 1 @ q2) =>
(fun (q3 : wlx1 @ 1 = 1 @ x1)
(s3 : (ehlnat_x1 [I] urnat_x1) @ 1 = 1 @ q3) =>
(fun (q4 : wlx0 @ 1 = 1 @ x0)
(s4 : (ehlnat_x0 [I] urnat_x0) @ 1 = 1 @ q4)
=>
(fun (b0 : wlx1 @ 1 = 1 @ wlx2)
(x : (... => ...) b0) =>
(fun (b1 : ...) (x3 : ...) => (...) (...) (...))
((...^ @ ...) @ concat_p1 (...))
(equiv_helper (ehrnat_y1 [-] ehlnat_x1)
(... @ ...) (... @ ...) 1 ))
(((concat_p1 (...))^ @ ((...) @ (...)^)) @
concat_p1 (1 @ wlx2))
(equiv_helper (ehrnat_z1 [-] ehlnat_x2)
((ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^)
(((...)^ @ (...)) @ concat_p1 (1 @ wlx2)) 1 ))
(ehlnat_x0 [I] urnat_x0)
(rlucancel 1 )) (ehlnat_x1 [I] urnat_x1)
(rlucancel 1 )) (ehlnat_x2 [I] urnat_x2)
(rlucancel 1 ))
((whiskerR 1 1 )^ @ (urnat_y0 [-] urnat_z0))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @ (urnat_y0 [-] urnat_z0))
(urnat_y0 [-] urnat_z0)
(whiskerR 1 1 ) 1 ))
((whiskerR 1 1 )^ @ (urnat_y1 [-] urnat_z1))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @ (urnat_y1 [-] urnat_z1))
(urnat_y1 [-] urnat_z1)
(whiskerR 1 1 ) 1 ))
((whiskerR 1 wrx2)^ @
((wlrnat_y_x [-] wlrnat_z_x) @ whiskerL wrx0 1 ))
(equiv_moveL_Mp
((whiskerR 1 wrx2)^ @
((wlrnat_y_x [-] wlrnat_z_x) @ whiskerL wrx0 1 ))
((wlrnat_y_x [-] wlrnat_z_x) @ whiskerL wrx0 1 )
(whiskerR 1 wrx2) 1 )
snapply equiv_path_ind_moveL_Mp. X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, x0, wrx0, wlx1, x1, wrx1, wlx2, x2, wrx2 : a = b wly0, y0 : b = b wly1, y1 : a = a wlz0, z0 : b = b wlz1, z1 : a = a urnat_x0 : wrx0 @ 1 = 1 @ x0 urnat_x1 : wrx1 @ 1 = 1 @ x1 urnat_x2 : wrx2 @ 1 = 1 @ x2 ulnat_y0 : wly0 @ 1 = 1 @ y0 urnat_y0 : 1 @ 1 = 1 @ y0ulnat_y1 : wly1 @ 1 = 1 @ y1 urnat_y1 : 1 @ 1 = 1 @ y1ulnat_z0 : wlz0 @ 1 = 1 @ z0 urnat_z0 : 1 @ 1 = 1 @ z0ulnat_z1 : wlz1 @ 1 = 1 @ z1 urnat_z1 : 1 @ 1 = 1 @ z1ulnat_yz0 : (wly0 @ wlz0) @ 1 = 1 @ (y0 @ z0) ulnat_yz1 : (wly1 @ wlz1) @ 1 = 1 @ (y1 @ z1) ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : 1 @ 1 = 1 @ wly0ehrnat_y1 : 1 @ 1 = 1 @ wly1ehrnat_z0 : 1 @ 1 = 1 @ wlz0ehrnat_z1 : 1 @ 1 = 1 @ wlz1ehrnat_yz1 : (1 @ 1 ) @ 1 = 1 @ (wly1 @ wlz1) wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0 H_ulnat_yz0 : ulnat_y0 [-] ulnat_z0 =
whiskerR 1 1 @ ulnat_yz0 H_ulnat_yz1 : ulnat_y1 [-] ulnat_z1 =
whiskerR 1 1 @ ulnat_yz1 H_ehrnat_yz1 : (ehrnat_y1 [-] ehrnat_z1) @
whiskerL 1 1 =
whiskerR 1 1 @ ehrnat_yz1 ehrnat_p1_y0 : (ehrnat_y0 [I] ulnat_y0) @ 1 =
1 @ urnat_y0 ehrnat_p1_y1 : (ehrnat_y1 [I] ulnat_y1) @ 1 =
1 @ urnat_y1 ehrnat_p1_z0 : (ehrnat_z0 [I] ulnat_z0) @ 1 =
1 @ urnat_z0 ehrnat_p1_z1 : (ehrnat_z1 [I] ulnat_z1) @ 1 =
1 @ urnat_z1
(fun (q : (1 @ 1 ) @ 1 = 1 @ (wly0 @ wlz0))
(s : (ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 1 =
whiskerR 1 1 @ q) =>
(fun
(q0 : (wly1 @ wlz1) @ wrx2 = wrx0 @ (wly0 @ wlz0))
(s0 : (wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 1 =
whiskerR 1 wrx2 @ q0) =>
(fun (q1 : (1 @ 1 ) @ 1 = 1 @ (y1 @ z1))
(s1 : urnat_y1 [-] urnat_z1 = whiskerR 1 1 @ q1)
=>
(fun (q2 : (1 @ 1 ) @ 1 = 1 @ (y0 @ z0))
(s2 : urnat_y0 [-] urnat_z0 = whiskerR 1 1 @ q2)
=>
(fun (q3 : wlx2 @ 1 = 1 @ x2)
(s3 : (ehlnat_x2 [I] urnat_x2) @ 1 = 1 @ q3) =>
(fun (q4 : wlx1 @ 1 = 1 @ x1)
(s4 : (ehlnat_x1 [I] urnat_x1) @ 1 = 1 @ q4)
=>
(fun (q5 : wlx0 @ 1 = 1 @ x0)
(s5 : (ehlnat_x0 [I] urnat_x0) @ 1 = 1 @ q5)
=>
(fun (b0 : ... = ...) (x : ... b0) =>
(... => ...) (... @ ...)
(equiv_helper ... ... ... 1 ))
(((...)^ @ (...)) @ concat_p1 (1 @ wlx2))
(equiv_helper (ehrnat_z1 [-] ehlnat_x2)
((...) @ whiskerL 1 wlrnat_z_x^)
((...) @ concat_p1 ...) 1 ))
(ehlnat_x0 [I] urnat_x0)
(rlucancel 1 )) (ehlnat_x1 [I] urnat_x1)
(rlucancel 1 )) (ehlnat_x2 [I] urnat_x2)
(rlucancel 1 ))
((whiskerR 1 1 )^ @ (urnat_y0 [-] urnat_z0))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @ (urnat_y0 [-] urnat_z0))
(urnat_y0 [-] urnat_z0)
(whiskerR 1 1 ) 1 ))
((whiskerR 1 1 )^ @ (urnat_y1 [-] urnat_z1))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @ (urnat_y1 [-] urnat_z1))
(urnat_y1 [-] urnat_z1)
(whiskerR 1 1 ) 1 ))
((whiskerR 1 wrx2)^ @
((wlrnat_y_x [-] wlrnat_z_x) @ whiskerL wrx0 1 ))
(equiv_moveL_Mp
((whiskerR 1 wrx2)^ @
((wlrnat_y_x [-] wlrnat_z_x) @ whiskerL wrx0 1 ))
((wlrnat_y_x [-] wlrnat_z_x) @ whiskerL wrx0 1 )
(whiskerR 1 wrx2) 1 ))
((whiskerR 1 1 )^ @
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 1 ))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 1 ))
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 1 )
(whiskerR 1 1 ) 1 )
revert ehrnat_yz1 H_ehrnat_yz1.X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, x0, wrx0, wlx1, x1, wrx1, wlx2, x2, wrx2 : a = b wly0, y0 : b = b wly1, y1 : a = a wlz0, z0 : b = b wlz1, z1 : a = a urnat_x0 : wrx0 @ 1 = 1 @ x0 urnat_x1 : wrx1 @ 1 = 1 @ x1 urnat_x2 : wrx2 @ 1 = 1 @ x2 ulnat_y0 : wly0 @ 1 = 1 @ y0 urnat_y0 : 1 @ 1 = 1 @ y0ulnat_y1 : wly1 @ 1 = 1 @ y1 urnat_y1 : 1 @ 1 = 1 @ y1ulnat_z0 : wlz0 @ 1 = 1 @ z0 urnat_z0 : 1 @ 1 = 1 @ z0ulnat_z1 : wlz1 @ 1 = 1 @ z1 urnat_z1 : 1 @ 1 = 1 @ z1ulnat_yz0 : (wly0 @ wlz0) @ 1 = 1 @ (y0 @ z0) ulnat_yz1 : (wly1 @ wlz1) @ 1 = 1 @ (y1 @ z1) ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : 1 @ 1 = 1 @ wly0ehrnat_y1 : 1 @ 1 = 1 @ wly1ehrnat_z0 : 1 @ 1 = 1 @ wlz0ehrnat_z1 : 1 @ 1 = 1 @ wlz1wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0 H_ulnat_yz0 : ulnat_y0 [-] ulnat_z0 =
whiskerR 1 1 @ ulnat_yz0 H_ulnat_yz1 : ulnat_y1 [-] ulnat_z1 =
whiskerR 1 1 @ ulnat_yz1 ehrnat_p1_y0 : (ehrnat_y0 [I] ulnat_y0) @ 1 =
1 @ urnat_y0 ehrnat_p1_y1 : (ehrnat_y1 [I] ulnat_y1) @ 1 =
1 @ urnat_y1 ehrnat_p1_z0 : (ehrnat_z0 [I] ulnat_z0) @ 1 =
1 @ urnat_z0 ehrnat_p1_z1 : (ehrnat_z1 [I] ulnat_z1) @ 1 =
1 @ urnat_z1
forall (ehrnat_yz1 : (1 @ 1 ) @ 1 = 1 @ (wly1 @ wlz1))
(H_ehrnat_yz1 : (ehrnat_y1 [-] ehrnat_z1) @
whiskerL 1 1 =
whiskerR 1 1 @ ehrnat_yz1),
(fun (q : (1 @ 1 ) @ 1 = 1 @ (wly0 @ wlz0))
(s : (ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 1 =
whiskerR 1 1 @ q) =>
(fun
(q0 : (wly1 @ wlz1) @ wrx2 = wrx0 @ (wly0 @ wlz0))
(s0 : (wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 1 =
whiskerR 1 wrx2 @ q0) =>
(fun (q1 : (1 @ 1 ) @ 1 = 1 @ (y1 @ z1))
(s1 : urnat_y1 [-] urnat_z1 = whiskerR 1 1 @ q1)
=>
(fun (q2 : (1 @ 1 ) @ 1 = 1 @ (y0 @ z0))
(s2 : urnat_y0 [-] urnat_z0 = whiskerR 1 1 @ q2)
=>
(fun (q3 : wlx2 @ 1 = 1 @ x2)
(s3 : (ehlnat_x2 [I] urnat_x2) @ 1 = 1 @ q3) =>
(fun (q4 : wlx1 @ 1 = 1 @ x1)
(s4 : (ehlnat_x1 [I] urnat_x1) @ 1 = 1 @ q4)
=>
(fun (q5 : wlx0 @ 1 = 1 @ x0)
(s5 : (...) @ 1 = 1 @ q5) =>
(fun (b0 : ...) (x : ...) => (...) (...) (...))
((...^ @ ...) @ concat_p1 (...))
(equiv_helper (ehrnat_z1 [-] ehlnat_x2)
(... @ ...) (... @ ...) 1 ))
(ehlnat_x0 [I] urnat_x0)
(rlucancel 1 )) (ehlnat_x1 [I] urnat_x1)
(rlucancel 1 )) (ehlnat_x2 [I] urnat_x2)
(rlucancel 1 ))
((whiskerR 1 1 )^ @ (urnat_y0 [-] urnat_z0))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @ (urnat_y0 [-] urnat_z0))
(urnat_y0 [-] urnat_z0)
(whiskerR 1 1 ) 1 ))
((whiskerR 1 1 )^ @ (urnat_y1 [-] urnat_z1))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @ (urnat_y1 [-] urnat_z1))
(urnat_y1 [-] urnat_z1)
(whiskerR 1 1 ) 1 ))
((whiskerR 1 wrx2)^ @
((wlrnat_y_x [-] wlrnat_z_x) @ whiskerL wrx0 1 ))
(equiv_moveL_Mp
((whiskerR 1 wrx2)^ @
((wlrnat_y_x [-] wlrnat_z_x) @ whiskerL wrx0 1 ))
((wlrnat_y_x [-] wlrnat_z_x) @ whiskerL wrx0 1 )
(whiskerR 1 wrx2) 1 ))
((whiskerR 1 1 )^ @
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 1 ))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 1 ))
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 1 )
(whiskerR 1 1 ) 1 )
snapply equiv_path_ind_moveL_Mp. X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, x0, wrx0, wlx1, x1, wrx1, wlx2, x2, wrx2 : a = b wly0, y0 : b = b wly1, y1 : a = a wlz0, z0 : b = b wlz1, z1 : a = a urnat_x0 : wrx0 @ 1 = 1 @ x0 urnat_x1 : wrx1 @ 1 = 1 @ x1 urnat_x2 : wrx2 @ 1 = 1 @ x2 ulnat_y0 : wly0 @ 1 = 1 @ y0 urnat_y0 : 1 @ 1 = 1 @ y0ulnat_y1 : wly1 @ 1 = 1 @ y1 urnat_y1 : 1 @ 1 = 1 @ y1ulnat_z0 : wlz0 @ 1 = 1 @ z0 urnat_z0 : 1 @ 1 = 1 @ z0ulnat_z1 : wlz1 @ 1 = 1 @ z1 urnat_z1 : 1 @ 1 = 1 @ z1ulnat_yz0 : (wly0 @ wlz0) @ 1 = 1 @ (y0 @ z0) ulnat_yz1 : (wly1 @ wlz1) @ 1 = 1 @ (y1 @ z1) ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : 1 @ 1 = 1 @ wly0ehrnat_y1 : 1 @ 1 = 1 @ wly1ehrnat_z0 : 1 @ 1 = 1 @ wlz0ehrnat_z1 : 1 @ 1 = 1 @ wlz1wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0 H_ulnat_yz0 : ulnat_y0 [-] ulnat_z0 =
whiskerR 1 1 @ ulnat_yz0 H_ulnat_yz1 : ulnat_y1 [-] ulnat_z1 =
whiskerR 1 1 @ ulnat_yz1 ehrnat_p1_y0 : (ehrnat_y0 [I] ulnat_y0) @ 1 =
1 @ urnat_y0 ehrnat_p1_y1 : (ehrnat_y1 [I] ulnat_y1) @ 1 =
1 @ urnat_y1 ehrnat_p1_z0 : (ehrnat_z0 [I] ulnat_z0) @ 1 =
1 @ urnat_z0 ehrnat_p1_z1 : (ehrnat_z1 [I] ulnat_z1) @ 1 =
1 @ urnat_z1
(fun (q : (1 @ 1 ) @ 1 = 1 @ (wly1 @ wlz1))
(s : (ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 1 =
whiskerR 1 1 @ q) =>
(fun (q0 : (1 @ 1 ) @ 1 = 1 @ (wly0 @ wlz0))
(s0 : (ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 1 =
whiskerR 1 1 @ q0) =>
(fun
(q1 : (wly1 @ wlz1) @ wrx2 = wrx0 @ (wly0 @ wlz0))
(s1 : (wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 1 =
whiskerR 1 wrx2 @ q1) =>
(fun (q2 : (1 @ 1 ) @ 1 = 1 @ (y1 @ z1))
(s2 : urnat_y1 [-] urnat_z1 = whiskerR 1 1 @ q2)
=>
(fun (q3 : (1 @ 1 ) @ 1 = 1 @ (y0 @ z0))
(s3 : urnat_y0 [-] urnat_z0 = whiskerR 1 1 @ q3)
=>
(fun (q4 : wlx2 @ 1 = 1 @ x2)
(s4 : (ehlnat_x2 [I] urnat_x2) @ 1 = 1 @ q4)
=>
(fun (q5 : wlx1 @ 1 = 1 @ x1)
(s5 : (ehlnat_x1 [I] urnat_x1) @ 1 = 1 @ q5)
=>
(fun (q6 : ... = ...) (s6 : ... = ...) =>
(... => ...) (... @ ...)
(equiv_helper ... ... ... 1 ))
(ehlnat_x0 [I] urnat_x0)
(rlucancel 1 )) (ehlnat_x1 [I] urnat_x1)
(rlucancel 1 )) (ehlnat_x2 [I] urnat_x2)
(rlucancel 1 ))
((whiskerR 1 1 )^ @ (urnat_y0 [-] urnat_z0))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @ (urnat_y0 [-] urnat_z0))
(urnat_y0 [-] urnat_z0)
(whiskerR 1 1 ) 1 ))
((whiskerR 1 1 )^ @ (urnat_y1 [-] urnat_z1))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @ (urnat_y1 [-] urnat_z1))
(urnat_y1 [-] urnat_z1)
(whiskerR 1 1 ) 1 ))
((whiskerR 1 wrx2)^ @
((wlrnat_y_x [-] wlrnat_z_x) @ whiskerL wrx0 1 ))
(equiv_moveL_Mp
((whiskerR 1 wrx2)^ @
((wlrnat_y_x [-] wlrnat_z_x) @ whiskerL wrx0 1 ))
((wlrnat_y_x [-] wlrnat_z_x) @ whiskerL wrx0 1 )
(whiskerR 1 wrx2) 1 ))
((whiskerR 1 1 )^ @
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 1 ))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 1 ))
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 1 )
(whiskerR 1 1 ) 1 ))
((whiskerR 1 1 )^ @
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 1 ))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 1 ))
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 1 )
(whiskerR 1 1 ) 1 )
revert ulnat_yz1 H_ulnat_yz1.X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, x0, wrx0, wlx1, x1, wrx1, wlx2, x2, wrx2 : a = b wly0, y0 : b = b wly1, y1 : a = a wlz0, z0 : b = b wlz1, z1 : a = a urnat_x0 : wrx0 @ 1 = 1 @ x0 urnat_x1 : wrx1 @ 1 = 1 @ x1 urnat_x2 : wrx2 @ 1 = 1 @ x2 ulnat_y0 : wly0 @ 1 = 1 @ y0 urnat_y0 : 1 @ 1 = 1 @ y0ulnat_y1 : wly1 @ 1 = 1 @ y1 urnat_y1 : 1 @ 1 = 1 @ y1ulnat_z0 : wlz0 @ 1 = 1 @ z0 urnat_z0 : 1 @ 1 = 1 @ z0ulnat_z1 : wlz1 @ 1 = 1 @ z1 urnat_z1 : 1 @ 1 = 1 @ z1ulnat_yz0 : (wly0 @ wlz0) @ 1 = 1 @ (y0 @ z0) ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : 1 @ 1 = 1 @ wly0ehrnat_y1 : 1 @ 1 = 1 @ wly1ehrnat_z0 : 1 @ 1 = 1 @ wlz0ehrnat_z1 : 1 @ 1 = 1 @ wlz1wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0 H_ulnat_yz0 : ulnat_y0 [-] ulnat_z0 =
whiskerR 1 1 @ ulnat_yz0 ehrnat_p1_y0 : (ehrnat_y0 [I] ulnat_y0) @ 1 =
1 @ urnat_y0 ehrnat_p1_y1 : (ehrnat_y1 [I] ulnat_y1) @ 1 =
1 @ urnat_y1 ehrnat_p1_z0 : (ehrnat_z0 [I] ulnat_z0) @ 1 =
1 @ urnat_z0 ehrnat_p1_z1 : (ehrnat_z1 [I] ulnat_z1) @ 1 =
1 @ urnat_z1
forall (ulnat_yz1 : (wly1 @ wlz1) @ 1 = 1 @ (y1 @ z1))
(H_ulnat_yz1 : ulnat_y1 [-] ulnat_z1 =
whiskerR 1 1 @ ulnat_yz1),
(fun (q : (1 @ 1 ) @ 1 = 1 @ (wly1 @ wlz1))
(s : (ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 1 =
whiskerR 1 1 @ q) =>
(fun (q0 : (1 @ 1 ) @ 1 = 1 @ (wly0 @ wlz0))
(s0 : (ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 1 =
whiskerR 1 1 @ q0) =>
(fun
(q1 : (wly1 @ wlz1) @ wrx2 = wrx0 @ (wly0 @ wlz0))
(s1 : (wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 1 =
whiskerR 1 wrx2 @ q1) =>
(fun (q2 : (1 @ 1 ) @ 1 = 1 @ (y1 @ z1))
(s2 : urnat_y1 [-] urnat_z1 = whiskerR 1 1 @ q2)
=>
(fun (q3 : (1 @ 1 ) @ 1 = 1 @ (y0 @ z0))
(s3 : urnat_y0 [-] urnat_z0 = whiskerR 1 1 @ q3)
=>
(fun (q4 : wlx2 @ 1 = 1 @ x2)
(s4 : (ehlnat_x2 [I] urnat_x2) @ 1 = 1 @ q4)
=>
(fun (q5 : wlx1 @ 1 = 1 @ x1)
(s5 : (...) @ 1 = 1 @ q5) =>
(fun (q6 : ...) (s6 : ...) => (...) (...) (...))
(ehlnat_x0 [I] urnat_x0)
(rlucancel 1 )) (ehlnat_x1 [I] urnat_x1)
(rlucancel 1 )) (ehlnat_x2 [I] urnat_x2)
(rlucancel 1 ))
((whiskerR 1 1 )^ @ (urnat_y0 [-] urnat_z0))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @ (urnat_y0 [-] urnat_z0))
(urnat_y0 [-] urnat_z0)
(whiskerR 1 1 ) 1 ))
((whiskerR 1 1 )^ @ (urnat_y1 [-] urnat_z1))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @ (urnat_y1 [-] urnat_z1))
(urnat_y1 [-] urnat_z1)
(whiskerR 1 1 ) 1 ))
((whiskerR 1 wrx2)^ @
((wlrnat_y_x [-] wlrnat_z_x) @ whiskerL wrx0 1 ))
(equiv_moveL_Mp
((whiskerR 1 wrx2)^ @
((wlrnat_y_x [-] wlrnat_z_x) @ whiskerL wrx0 1 ))
((wlrnat_y_x [-] wlrnat_z_x) @ whiskerL wrx0 1 )
(whiskerR 1 wrx2) 1 ))
((whiskerR 1 1 )^ @
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 1 ))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 1 ))
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 1 )
(whiskerR 1 1 ) 1 ))
((whiskerR 1 1 )^ @
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 1 ))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 1 ))
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 1 )
(whiskerR 1 1 ) 1 )
snapply equiv_path_ind_moveL_Mp. X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, x0, wrx0, wlx1, x1, wrx1, wlx2, x2, wrx2 : a = b wly0, y0 : b = b wly1, y1 : a = a wlz0, z0 : b = b wlz1, z1 : a = a urnat_x0 : wrx0 @ 1 = 1 @ x0 urnat_x1 : wrx1 @ 1 = 1 @ x1 urnat_x2 : wrx2 @ 1 = 1 @ x2 ulnat_y0 : wly0 @ 1 = 1 @ y0 urnat_y0 : 1 @ 1 = 1 @ y0ulnat_y1 : wly1 @ 1 = 1 @ y1 urnat_y1 : 1 @ 1 = 1 @ y1ulnat_z0 : wlz0 @ 1 = 1 @ z0 urnat_z0 : 1 @ 1 = 1 @ z0ulnat_z1 : wlz1 @ 1 = 1 @ z1 urnat_z1 : 1 @ 1 = 1 @ z1ulnat_yz0 : (wly0 @ wlz0) @ 1 = 1 @ (y0 @ z0) ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : 1 @ 1 = 1 @ wly0ehrnat_y1 : 1 @ 1 = 1 @ wly1ehrnat_z0 : 1 @ 1 = 1 @ wlz0ehrnat_z1 : 1 @ 1 = 1 @ wlz1wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0 H_ulnat_yz0 : ulnat_y0 [-] ulnat_z0 =
whiskerR 1 1 @ ulnat_yz0 ehrnat_p1_y0 : (ehrnat_y0 [I] ulnat_y0) @ 1 =
1 @ urnat_y0 ehrnat_p1_y1 : (ehrnat_y1 [I] ulnat_y1) @ 1 =
1 @ urnat_y1 ehrnat_p1_z0 : (ehrnat_z0 [I] ulnat_z0) @ 1 =
1 @ urnat_z0 ehrnat_p1_z1 : (ehrnat_z1 [I] ulnat_z1) @ 1 =
1 @ urnat_z1
(fun (q : (wly1 @ wlz1) @ 1 = 1 @ (y1 @ z1))
(s : ulnat_y1 [-] ulnat_z1 = whiskerR 1 1 @ q) =>
(fun (q0 : (1 @ 1 ) @ 1 = 1 @ (wly1 @ wlz1))
(s0 : (ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 1 =
whiskerR 1 1 @ q0) =>
(fun (q1 : (1 @ 1 ) @ 1 = 1 @ (wly0 @ wlz0))
(s1 : (ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 1 =
whiskerR 1 1 @ q1) =>
(fun
(q2 : (wly1 @ wlz1) @ wrx2 =
wrx0 @ (wly0 @ wlz0))
(s2 : (wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 1 =
whiskerR 1 wrx2 @ q2) =>
(fun (q3 : (1 @ 1 ) @ 1 = 1 @ (y1 @ z1))
(s3 : urnat_y1 [-] urnat_z1 = whiskerR 1 1 @ q3)
=>
(fun (q4 : (1 @ 1 ) @ 1 = 1 @ (y0 @ z0))
(s4 : urnat_y0 [-] urnat_z0 =
whiskerR 1 1 @ q4) =>
(fun (q5 : wlx2 @ 1 = 1 @ x2)
(s5 : (ehlnat_x2 [I] urnat_x2) @ 1 = 1 @ q5)
=>
(fun (q6 : ... = ...) (s6 : ... = ...) =>
(... => ...) (ehlnat_x0 [I] urnat_x0)
(rlucancel 1 ))
(ehlnat_x1 [I] urnat_x1)
(rlucancel 1 )) (ehlnat_x2 [I] urnat_x2)
(rlucancel 1 ))
((whiskerR 1 1 )^ @ (urnat_y0 [-] urnat_z0))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @ (urnat_y0 [-] urnat_z0))
(urnat_y0 [-] urnat_z0)
(whiskerR 1 1 ) 1 ))
((whiskerR 1 1 )^ @ (urnat_y1 [-] urnat_z1))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @ (urnat_y1 [-] urnat_z1))
(urnat_y1 [-] urnat_z1)
(whiskerR 1 1 ) 1 ))
((whiskerR 1 wrx2)^ @
((wlrnat_y_x [-] wlrnat_z_x) @ whiskerL wrx0 1 ))
(equiv_moveL_Mp
((whiskerR 1 wrx2)^ @
((wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 1 ))
((wlrnat_y_x [-] wlrnat_z_x) @ whiskerL wrx0 1 )
(whiskerR 1 wrx2) 1 ))
((whiskerR 1 1 )^ @
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 1 ))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 1 ))
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 1 )
(whiskerR 1 1 ) 1 ))
((whiskerR 1 1 )^ @
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 1 ))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 1 ))
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 1 )
(whiskerR 1 1 ) 1 ))
((whiskerR 1 1 )^ @ (ulnat_y1 [-] ulnat_z1))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @ (ulnat_y1 [-] ulnat_z1))
(ulnat_y1 [-] ulnat_z1)
(whiskerR 1 1 ) 1 )
revert ulnat_yz0 H_ulnat_yz0.X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, x0, wrx0, wlx1, x1, wrx1, wlx2, x2, wrx2 : a = b wly0, y0 : b = b wly1, y1 : a = a wlz0, z0 : b = b wlz1, z1 : a = a urnat_x0 : wrx0 @ 1 = 1 @ x0 urnat_x1 : wrx1 @ 1 = 1 @ x1 urnat_x2 : wrx2 @ 1 = 1 @ x2 ulnat_y0 : wly0 @ 1 = 1 @ y0 urnat_y0 : 1 @ 1 = 1 @ y0ulnat_y1 : wly1 @ 1 = 1 @ y1 urnat_y1 : 1 @ 1 = 1 @ y1ulnat_z0 : wlz0 @ 1 = 1 @ z0 urnat_z0 : 1 @ 1 = 1 @ z0ulnat_z1 : wlz1 @ 1 = 1 @ z1 urnat_z1 : 1 @ 1 = 1 @ z1ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : 1 @ 1 = 1 @ wly0ehrnat_y1 : 1 @ 1 = 1 @ wly1ehrnat_z0 : 1 @ 1 = 1 @ wlz0ehrnat_z1 : 1 @ 1 = 1 @ wlz1wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0 ehrnat_p1_y0 : (ehrnat_y0 [I] ulnat_y0) @ 1 =
1 @ urnat_y0 ehrnat_p1_y1 : (ehrnat_y1 [I] ulnat_y1) @ 1 =
1 @ urnat_y1 ehrnat_p1_z0 : (ehrnat_z0 [I] ulnat_z0) @ 1 =
1 @ urnat_z0 ehrnat_p1_z1 : (ehrnat_z1 [I] ulnat_z1) @ 1 =
1 @ urnat_z1
forall (ulnat_yz0 : (wly0 @ wlz0) @ 1 = 1 @ (y0 @ z0))
(H_ulnat_yz0 : ulnat_y0 [-] ulnat_z0 =
whiskerR 1 1 @ ulnat_yz0),
(fun (q : (wly1 @ wlz1) @ 1 = 1 @ (y1 @ z1))
(s : ulnat_y1 [-] ulnat_z1 = whiskerR 1 1 @ q) =>
(fun (q0 : (1 @ 1 ) @ 1 = 1 @ (wly1 @ wlz1))
(s0 : (ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 1 =
whiskerR 1 1 @ q0) =>
(fun (q1 : (1 @ 1 ) @ 1 = 1 @ (wly0 @ wlz0))
(s1 : (ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 1 =
whiskerR 1 1 @ q1) =>
(fun
(q2 : (wly1 @ wlz1) @ wrx2 =
wrx0 @ (wly0 @ wlz0))
(s2 : (wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 1 =
whiskerR 1 wrx2 @ q2) =>
(fun (q3 : (1 @ 1 ) @ 1 = 1 @ (y1 @ z1))
(s3 : urnat_y1 [-] urnat_z1 = whiskerR 1 1 @ q3)
=>
(fun (q4 : (1 @ 1 ) @ 1 = 1 @ (y0 @ z0))
(s4 : urnat_y0 [-] urnat_z0 =
whiskerR 1 1 @ q4) =>
(fun (q5 : wlx2 @ 1 = 1 @ x2)
(s5 : (...) @ 1 = 1 @ q5) =>
(fun (q6 : ...) (s6 : ...) => (...) (...) (...))
(ehlnat_x1 [I] urnat_x1)
(rlucancel 1 )) (ehlnat_x2 [I] urnat_x2)
(rlucancel 1 ))
((whiskerR 1 1 )^ @ (urnat_y0 [-] urnat_z0))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @ (urnat_y0 [-] urnat_z0))
(urnat_y0 [-] urnat_z0)
(whiskerR 1 1 ) 1 ))
((whiskerR 1 1 )^ @ (urnat_y1 [-] urnat_z1))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @ (urnat_y1 [-] urnat_z1))
(urnat_y1 [-] urnat_z1)
(whiskerR 1 1 ) 1 ))
((whiskerR 1 wrx2)^ @
((wlrnat_y_x [-] wlrnat_z_x) @ whiskerL wrx0 1 ))
(equiv_moveL_Mp
((whiskerR 1 wrx2)^ @
((wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 1 ))
((wlrnat_y_x [-] wlrnat_z_x) @ whiskerL wrx0 1 )
(whiskerR 1 wrx2) 1 ))
((whiskerR 1 1 )^ @
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 1 ))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 1 ))
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 1 )
(whiskerR 1 1 ) 1 ))
((whiskerR 1 1 )^ @
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 1 ))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 1 ))
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 1 )
(whiskerR 1 1 ) 1 ))
((whiskerR 1 1 )^ @ (ulnat_y1 [-] ulnat_z1))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @ (ulnat_y1 [-] ulnat_z1))
(ulnat_y1 [-] ulnat_z1)
(whiskerR 1 1 ) 1 )
snapply equiv_path_ind_moveL_Mp. X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, x0, wrx0, wlx1, x1, wrx1, wlx2, x2, wrx2 : a = b wly0, y0 : b = b wly1, y1 : a = a wlz0, z0 : b = b wlz1, z1 : a = a urnat_x0 : wrx0 @ 1 = 1 @ x0 urnat_x1 : wrx1 @ 1 = 1 @ x1 urnat_x2 : wrx2 @ 1 = 1 @ x2 ulnat_y0 : wly0 @ 1 = 1 @ y0 urnat_y0 : 1 @ 1 = 1 @ y0ulnat_y1 : wly1 @ 1 = 1 @ y1 urnat_y1 : 1 @ 1 = 1 @ y1ulnat_z0 : wlz0 @ 1 = 1 @ z0 urnat_z0 : 1 @ 1 = 1 @ z0ulnat_z1 : wlz1 @ 1 = 1 @ z1 urnat_z1 : 1 @ 1 = 1 @ z1ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : 1 @ 1 = 1 @ wly0ehrnat_y1 : 1 @ 1 = 1 @ wly1ehrnat_z0 : 1 @ 1 = 1 @ wlz0ehrnat_z1 : 1 @ 1 = 1 @ wlz1wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0 ehrnat_p1_y0 : (ehrnat_y0 [I] ulnat_y0) @ 1 =
1 @ urnat_y0 ehrnat_p1_y1 : (ehrnat_y1 [I] ulnat_y1) @ 1 =
1 @ urnat_y1 ehrnat_p1_z0 : (ehrnat_z0 [I] ulnat_z0) @ 1 =
1 @ urnat_z0 ehrnat_p1_z1 : (ehrnat_z1 [I] ulnat_z1) @ 1 =
1 @ urnat_z1
(fun (q : (wly0 @ wlz0) @ 1 = 1 @ (y0 @ z0))
(s : ulnat_y0 [-] ulnat_z0 = whiskerR 1 1 @ q) =>
(fun (q0 : (wly1 @ wlz1) @ 1 = 1 @ (y1 @ z1))
(s0 : ulnat_y1 [-] ulnat_z1 = whiskerR 1 1 @ q0)
=>
(fun (q1 : (1 @ 1 ) @ 1 = 1 @ (wly1 @ wlz1))
(s1 : (ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 1 =
whiskerR 1 1 @ q1) =>
(fun (q2 : (1 @ 1 ) @ 1 = 1 @ (wly0 @ wlz0))
(s2 : (ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 1 =
whiskerR 1 1 @ q2) =>
(fun
(q3 : (wly1 @ wlz1) @ wrx2 =
wrx0 @ (wly0 @ wlz0))
(s3 : (wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 1 =
whiskerR 1 wrx2 @ q3) =>
(fun (q4 : (1 @ 1 ) @ 1 = 1 @ (y1 @ z1))
(s4 : urnat_y1 [-] urnat_z1 =
whiskerR 1 1 @ q4) =>
(fun (q5 : (1 @ 1 ) @ 1 = 1 @ (y0 @ z0))
(s5 : urnat_y0 [-] urnat_z0 =
whiskerR 1 1 @ q5) =>
(fun (q6 : ... = ...) (s6 : ... = ...) =>
(... => ...) (ehlnat_x1 [I] urnat_x1)
(rlucancel 1 ))
(ehlnat_x2 [I] urnat_x2)
(rlucancel 1 ))
((whiskerR 1 1 )^ @ (urnat_y0 [-] urnat_z0))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @ (urnat_y0 [-] urnat_z0))
(urnat_y0 [-] urnat_z0)
(whiskerR 1 1 ) 1 ))
((whiskerR 1 1 )^ @ (urnat_y1 [-] urnat_z1))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @ (urnat_y1 [-] urnat_z1))
(urnat_y1 [-] urnat_z1)
(whiskerR 1 1 ) 1 ))
((whiskerR 1 wrx2)^ @
((wlrnat_y_x [-] wlrnat_z_x) @ whiskerL wrx0 1 ))
(equiv_moveL_Mp
((whiskerR 1 wrx2)^ @
((wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 1 ))
((wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 1 )
(whiskerR 1 wrx2) 1 ))
((whiskerR 1 1 )^ @
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 1 ))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 1 ))
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 1 )
(whiskerR 1 1 ) 1 ))
((whiskerR 1 1 )^ @
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 1 ))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 1 ))
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 1 )
(whiskerR 1 1 ) 1 ))
((whiskerR 1 1 )^ @ (ulnat_y1 [-] ulnat_z1))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @ (ulnat_y1 [-] ulnat_z1))
(ulnat_y1 [-] ulnat_z1)
(whiskerR 1 1 ) 1 ))
((whiskerR 1 1 )^ @ (ulnat_y0 [-] ulnat_z0))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @ (ulnat_y0 [-] ulnat_z0))
(ulnat_y0 [-] ulnat_z0)
(whiskerR 1 1 ) 1 )
revert urnat_y0 ehrnat_p1_y0.X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, x0, wrx0, wlx1, x1, wrx1, wlx2, x2, wrx2 : a = b wly0, y0 : b = b wly1, y1 : a = a wlz0, z0 : b = b wlz1, z1 : a = a urnat_x0 : wrx0 @ 1 = 1 @ x0 urnat_x1 : wrx1 @ 1 = 1 @ x1 urnat_x2 : wrx2 @ 1 = 1 @ x2 ulnat_y0 : wly0 @ 1 = 1 @ y0 ulnat_y1 : wly1 @ 1 = 1 @ y1 urnat_y1 : 1 @ 1 = 1 @ y1ulnat_z0 : wlz0 @ 1 = 1 @ z0 urnat_z0 : 1 @ 1 = 1 @ z0ulnat_z1 : wlz1 @ 1 = 1 @ z1 urnat_z1 : 1 @ 1 = 1 @ z1ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : 1 @ 1 = 1 @ wly0ehrnat_y1 : 1 @ 1 = 1 @ wly1ehrnat_z0 : 1 @ 1 = 1 @ wlz0ehrnat_z1 : 1 @ 1 = 1 @ wlz1wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0 ehrnat_p1_y1 : (ehrnat_y1 [I] ulnat_y1) @ 1 =
1 @ urnat_y1 ehrnat_p1_z0 : (ehrnat_z0 [I] ulnat_z0) @ 1 =
1 @ urnat_z0 ehrnat_p1_z1 : (ehrnat_z1 [I] ulnat_z1) @ 1 =
1 @ urnat_z1
forall (urnat_y0 : 1 @ 1 = 1 @ y0)
(ehrnat_p1_y0 : (ehrnat_y0 [I] ulnat_y0) @ 1 =
1 @ urnat_y0),
(fun (q : (wly0 @ wlz0) @ 1 = 1 @ (y0 @ z0))
(s : ulnat_y0 [-] ulnat_z0 = whiskerR 1 1 @ q) =>
(fun (q0 : (wly1 @ wlz1) @ 1 = 1 @ (y1 @ z1))
(s0 : ulnat_y1 [-] ulnat_z1 = whiskerR 1 1 @ q0)
=>
(fun (q1 : (1 @ 1 ) @ 1 = 1 @ (wly1 @ wlz1))
(s1 : (ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 1 =
whiskerR 1 1 @ q1) =>
(fun (q2 : (1 @ 1 ) @ 1 = 1 @ (wly0 @ wlz0))
(s2 : (ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 1 =
whiskerR 1 1 @ q2) =>
(fun
(q3 : (wly1 @ wlz1) @ wrx2 =
wrx0 @ (wly0 @ wlz0))
(s3 : (wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 1 =
whiskerR 1 wrx2 @ q3) =>
(fun (q4 : (1 @ 1 ) @ 1 = 1 @ (y1 @ z1))
(s4 : urnat_y1 [-] urnat_z1 =
whiskerR 1 1 @ q4) =>
(fun (q5 : (...) @ 1 = 1 @ (...))
(s5 : urnat_y0 [-] urnat_z0 =
whiskerR 1 1 @ q5) =>
(fun (q6 : ...) (s6 : ...) => (...) (...) (...))
(ehlnat_x2 [I] urnat_x2)
(rlucancel 1 ))
((whiskerR 1 1 )^ @ (urnat_y0 [-] urnat_z0))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @ (urnat_y0 [-] urnat_z0))
(urnat_y0 [-] urnat_z0)
(whiskerR 1 1 ) 1 ))
((whiskerR 1 1 )^ @ (urnat_y1 [-] urnat_z1))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @ (urnat_y1 [-] urnat_z1))
(urnat_y1 [-] urnat_z1)
(whiskerR 1 1 ) 1 ))
((whiskerR 1 wrx2)^ @
((wlrnat_y_x [-] wlrnat_z_x) @ whiskerL wrx0 1 ))
(equiv_moveL_Mp
((whiskerR 1 wrx2)^ @
((wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 1 ))
((wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 1 )
(whiskerR 1 wrx2) 1 ))
((whiskerR 1 1 )^ @
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 1 ))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 1 ))
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 1 )
(whiskerR 1 1 ) 1 ))
((whiskerR 1 1 )^ @
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 1 ))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 1 ))
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 1 )
(whiskerR 1 1 ) 1 ))
((whiskerR 1 1 )^ @ (ulnat_y1 [-] ulnat_z1))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @ (ulnat_y1 [-] ulnat_z1))
(ulnat_y1 [-] ulnat_z1)
(whiskerR 1 1 ) 1 ))
((whiskerR 1 1 )^ @ (ulnat_y0 [-] ulnat_z0))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @ (ulnat_y0 [-] ulnat_z0))
(ulnat_y0 [-] ulnat_z0)
(whiskerR 1 1 ) 1 )
snapply equiv_path_ind_rlucancel. X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, x0, wrx0, wlx1, x1, wrx1, wlx2, x2, wrx2 : a = b wly0, y0 : b = b wly1, y1 : a = a wlz0, z0 : b = b wlz1, z1 : a = a urnat_x0 : wrx0 @ 1 = 1 @ x0 urnat_x1 : wrx1 @ 1 = 1 @ x1 urnat_x2 : wrx2 @ 1 = 1 @ x2 ulnat_y0 : wly0 @ 1 = 1 @ y0 ulnat_y1 : wly1 @ 1 = 1 @ y1 urnat_y1 : 1 @ 1 = 1 @ y1ulnat_z0 : wlz0 @ 1 = 1 @ z0 urnat_z0 : 1 @ 1 = 1 @ z0ulnat_z1 : wlz1 @ 1 = 1 @ z1 urnat_z1 : 1 @ 1 = 1 @ z1ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : 1 @ 1 = 1 @ wly0ehrnat_y1 : 1 @ 1 = 1 @ wly1ehrnat_z0 : 1 @ 1 = 1 @ wlz0ehrnat_z1 : 1 @ 1 = 1 @ wlz1wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0 ehrnat_p1_y1 : (ehrnat_y1 [I] ulnat_y1) @ 1 =
1 @ urnat_y1 ehrnat_p1_z0 : (ehrnat_z0 [I] ulnat_z0) @ 1 =
1 @ urnat_z0 ehrnat_p1_z1 : (ehrnat_z1 [I] ulnat_z1) @ 1 =
1 @ urnat_z1
(fun (q : 1 @ 1 = 1 @ y0)
(s : (ehrnat_y0 [I] ulnat_y0) @ 1 = 1 @ q) =>
(fun (q0 : (wly0 @ wlz0) @ 1 = 1 @ (y0 @ z0))
(s0 : ulnat_y0 [-] ulnat_z0 = whiskerR 1 1 @ q0)
=>
(fun (q1 : (wly1 @ wlz1) @ 1 = 1 @ (y1 @ z1))
(s1 : ulnat_y1 [-] ulnat_z1 = whiskerR 1 1 @ q1)
=>
(fun (q2 : (1 @ 1 ) @ 1 = 1 @ (wly1 @ wlz1))
(s2 : (ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 1 =
whiskerR 1 1 @ q2) =>
(fun (q3 : (1 @ 1 ) @ 1 = 1 @ (wly0 @ wlz0))
(s3 : (ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 1 =
whiskerR 1 1 @ q3) =>
(fun
(q4 : (wly1 @ wlz1) @ wrx2 =
wrx0 @ (wly0 @ wlz0))
(s4 : (wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 1 =
whiskerR 1 wrx2 @ q4) =>
(fun (q5 : (1 @ 1 ) @ 1 = 1 @ (y1 @ z1))
(s5 : urnat_y1 [-] urnat_z1 =
whiskerR 1 1 @ q5) =>
(fun (q6 : ... = ...) (s6 : ... = ...) =>
(... => ...) (ehlnat_x2 [I] urnat_x2)
(rlucancel 1 ))
((whiskerR 1 1 )^ @ (q [-] urnat_z0))
(equiv_moveL_Mp
((...)^ @ (...))
(q [-] urnat_z0)
(whiskerR 1 1 ) 1 ))
((whiskerR 1 1 )^ @ (urnat_y1 [-] urnat_z1))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @ (urnat_y1 [-] urnat_z1))
(urnat_y1 [-] urnat_z1)
(whiskerR 1 1 ) 1 ))
((whiskerR 1 wrx2)^ @
((wlrnat_y_x [-] wlrnat_z_x) @ whiskerL wrx0 1 ))
(equiv_moveL_Mp
((whiskerR 1 wrx2)^ @
((wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 1 ))
((wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 1 )
(whiskerR 1 wrx2) 1 ))
((whiskerR 1 1 )^ @
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 1 ))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 1 ))
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 1 )
(whiskerR 1 1 ) 1 ))
((whiskerR 1 1 )^ @
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 1 ))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 1 ))
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 1 )
(whiskerR 1 1 ) 1 ))
((whiskerR 1 1 )^ @ (ulnat_y1 [-] ulnat_z1))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @ (ulnat_y1 [-] ulnat_z1))
(ulnat_y1 [-] ulnat_z1)
(whiskerR 1 1 ) 1 ))
((whiskerR 1 1 )^ @ (ulnat_y0 [-] ulnat_z0))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @ (ulnat_y0 [-] ulnat_z0))
(ulnat_y0 [-] ulnat_z0)
(whiskerR 1 1 ) 1 ))
(ehrnat_y0 [I] ulnat_y0)
(rlucancel 1 )
revert urnat_y1 ehrnat_p1_y1.X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, x0, wrx0, wlx1, x1, wrx1, wlx2, x2, wrx2 : a = b wly0, y0 : b = b wly1, y1 : a = a wlz0, z0 : b = b wlz1, z1 : a = a urnat_x0 : wrx0 @ 1 = 1 @ x0 urnat_x1 : wrx1 @ 1 = 1 @ x1 urnat_x2 : wrx2 @ 1 = 1 @ x2 ulnat_y0 : wly0 @ 1 = 1 @ y0 ulnat_y1 : wly1 @ 1 = 1 @ y1 ulnat_z0 : wlz0 @ 1 = 1 @ z0 urnat_z0 : 1 @ 1 = 1 @ z0ulnat_z1 : wlz1 @ 1 = 1 @ z1 urnat_z1 : 1 @ 1 = 1 @ z1ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : 1 @ 1 = 1 @ wly0ehrnat_y1 : 1 @ 1 = 1 @ wly1ehrnat_z0 : 1 @ 1 = 1 @ wlz0ehrnat_z1 : 1 @ 1 = 1 @ wlz1wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0 ehrnat_p1_z0 : (ehrnat_z0 [I] ulnat_z0) @ 1 =
1 @ urnat_z0 ehrnat_p1_z1 : (ehrnat_z1 [I] ulnat_z1) @ 1 =
1 @ urnat_z1
forall (urnat_y1 : 1 @ 1 = 1 @ y1)
(ehrnat_p1_y1 : (ehrnat_y1 [I] ulnat_y1) @ 1 =
1 @ urnat_y1),
(fun (q : 1 @ 1 = 1 @ y0)
(s : (ehrnat_y0 [I] ulnat_y0) @ 1 = 1 @ q) =>
(fun (q0 : (wly0 @ wlz0) @ 1 = 1 @ (y0 @ z0))
(s0 : ulnat_y0 [-] ulnat_z0 = whiskerR 1 1 @ q0)
=>
(fun (q1 : (wly1 @ wlz1) @ 1 = 1 @ (y1 @ z1))
(s1 : ulnat_y1 [-] ulnat_z1 = whiskerR 1 1 @ q1)
=>
(fun (q2 : (1 @ 1 ) @ 1 = 1 @ (wly1 @ wlz1))
(s2 : (ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 1 =
whiskerR 1 1 @ q2) =>
(fun (q3 : (1 @ 1 ) @ 1 = 1 @ (wly0 @ wlz0))
(s3 : (ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 1 =
whiskerR 1 1 @ q3) =>
(fun
(q4 : (wly1 @ wlz1) @ wrx2 =
wrx0 @ (wly0 @ wlz0))
(s4 : (wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 1 =
whiskerR 1 wrx2 @ q4) =>
(fun (q5 : (...) @ 1 = 1 @ (...))
(s5 : urnat_y1 [-] urnat_z1 =
whiskerR 1 1 @ q5) =>
(fun (q6 : ...) (s6 : ...) => (...) (...) (...))
((whiskerR 1 1 )^ @ (q [-] urnat_z0))
(equiv_moveL_Mp
(...^ @ ...)
(q [-] urnat_z0)
(whiskerR 1 1 ) 1 ))
((whiskerR 1 1 )^ @ (urnat_y1 [-] urnat_z1))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @ (urnat_y1 [-] urnat_z1))
(urnat_y1 [-] urnat_z1)
(whiskerR 1 1 ) 1 ))
((whiskerR 1 wrx2)^ @
((wlrnat_y_x [-] wlrnat_z_x) @ whiskerL wrx0 1 ))
(equiv_moveL_Mp
((whiskerR 1 wrx2)^ @
((wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 1 ))
((wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 1 )
(whiskerR 1 wrx2) 1 ))
((whiskerR 1 1 )^ @
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 1 ))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 1 ))
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 1 )
(whiskerR 1 1 ) 1 ))
((whiskerR 1 1 )^ @
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 1 ))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 1 ))
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 1 )
(whiskerR 1 1 ) 1 ))
((whiskerR 1 1 )^ @ (ulnat_y1 [-] ulnat_z1))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @ (ulnat_y1 [-] ulnat_z1))
(ulnat_y1 [-] ulnat_z1)
(whiskerR 1 1 ) 1 ))
((whiskerR 1 1 )^ @ (ulnat_y0 [-] ulnat_z0))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @ (ulnat_y0 [-] ulnat_z0))
(ulnat_y0 [-] ulnat_z0)
(whiskerR 1 1 ) 1 ))
(ehrnat_y0 [I] ulnat_y0)
(rlucancel 1 )
snapply equiv_path_ind_rlucancel. X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, x0, wrx0, wlx1, x1, wrx1, wlx2, x2, wrx2 : a = b wly0, y0 : b = b wly1, y1 : a = a wlz0, z0 : b = b wlz1, z1 : a = a urnat_x0 : wrx0 @ 1 = 1 @ x0 urnat_x1 : wrx1 @ 1 = 1 @ x1 urnat_x2 : wrx2 @ 1 = 1 @ x2 ulnat_y0 : wly0 @ 1 = 1 @ y0 ulnat_y1 : wly1 @ 1 = 1 @ y1 ulnat_z0 : wlz0 @ 1 = 1 @ z0 urnat_z0 : 1 @ 1 = 1 @ z0ulnat_z1 : wlz1 @ 1 = 1 @ z1 urnat_z1 : 1 @ 1 = 1 @ z1ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : 1 @ 1 = 1 @ wly0ehrnat_y1 : 1 @ 1 = 1 @ wly1ehrnat_z0 : 1 @ 1 = 1 @ wlz0ehrnat_z1 : 1 @ 1 = 1 @ wlz1wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0 ehrnat_p1_z0 : (ehrnat_z0 [I] ulnat_z0) @ 1 =
1 @ urnat_z0 ehrnat_p1_z1 : (ehrnat_z1 [I] ulnat_z1) @ 1 =
1 @ urnat_z1
(fun (q : 1 @ 1 = 1 @ y1)
(s : (ehrnat_y1 [I] ulnat_y1) @ 1 = 1 @ q) =>
(fun (q0 : 1 @ 1 = 1 @ y0)
(s0 : (ehrnat_y0 [I] ulnat_y0) @ 1 = 1 @ q0) =>
(fun (q1 : (wly0 @ wlz0) @ 1 = 1 @ (y0 @ z0))
(s1 : ulnat_y0 [-] ulnat_z0 = whiskerR 1 1 @ q1)
=>
(fun (q2 : (wly1 @ wlz1) @ 1 = 1 @ (y1 @ z1))
(s2 : ulnat_y1 [-] ulnat_z1 = whiskerR 1 1 @ q2)
=>
(fun (q3 : (1 @ 1 ) @ 1 = 1 @ (wly1 @ wlz1))
(s3 : (ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 1 =
whiskerR 1 1 @ q3) =>
(fun (q4 : (1 @ 1 ) @ 1 = 1 @ (wly0 @ wlz0))
(s4 : (ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 1 =
whiskerR 1 1 @ q4) =>
(fun
(q5 : (wly1 @ wlz1) @ wrx2 =
wrx0 @ (wly0 @ wlz0))
(s5 : (wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 1 =
whiskerR 1 wrx2 @ q5) =>
(fun (q6 : ... = ...) (s6 : ... = ...) =>
(... => ...) (...^ @ ...)
(equiv_moveL_Mp ... ... ... 1 ))
((whiskerR 1 1 )^ @ (q [-] urnat_z1))
(equiv_moveL_Mp
((...)^ @ (...))
(q [-] urnat_z1)
(whiskerR 1 1 ) 1 ))
((whiskerR 1 wrx2)^ @
((wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 1 ))
(equiv_moveL_Mp
((whiskerR 1 wrx2)^ @
((wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 1 ))
((wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 1 )
(whiskerR 1 wrx2) 1 ))
((whiskerR 1 1 )^ @
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 1 ))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 1 ))
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 1 )
(whiskerR 1 1 ) 1 ))
((whiskerR 1 1 )^ @
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 1 ))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 1 ))
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 1 )
(whiskerR 1 1 ) 1 ))
((whiskerR 1 1 )^ @ (ulnat_y1 [-] ulnat_z1))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @ (ulnat_y1 [-] ulnat_z1))
(ulnat_y1 [-] ulnat_z1)
(whiskerR 1 1 ) 1 ))
((whiskerR 1 1 )^ @ (ulnat_y0 [-] ulnat_z0))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @ (ulnat_y0 [-] ulnat_z0))
(ulnat_y0 [-] ulnat_z0)
(whiskerR 1 1 ) 1 ))
(ehrnat_y0 [I] ulnat_y0)
(rlucancel 1 )) (ehrnat_y1 [I] ulnat_y1)
(rlucancel 1 )
revert urnat_z0 ehrnat_p1_z0.X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, x0, wrx0, wlx1, x1, wrx1, wlx2, x2, wrx2 : a = b wly0, y0 : b = b wly1, y1 : a = a wlz0, z0 : b = b wlz1, z1 : a = a urnat_x0 : wrx0 @ 1 = 1 @ x0 urnat_x1 : wrx1 @ 1 = 1 @ x1 urnat_x2 : wrx2 @ 1 = 1 @ x2 ulnat_y0 : wly0 @ 1 = 1 @ y0 ulnat_y1 : wly1 @ 1 = 1 @ y1 ulnat_z0 : wlz0 @ 1 = 1 @ z0 ulnat_z1 : wlz1 @ 1 = 1 @ z1 urnat_z1 : 1 @ 1 = 1 @ z1ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : 1 @ 1 = 1 @ wly0ehrnat_y1 : 1 @ 1 = 1 @ wly1ehrnat_z0 : 1 @ 1 = 1 @ wlz0ehrnat_z1 : 1 @ 1 = 1 @ wlz1wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0 ehrnat_p1_z1 : (ehrnat_z1 [I] ulnat_z1) @ 1 =
1 @ urnat_z1
forall (urnat_z0 : 1 @ 1 = 1 @ z0)
(ehrnat_p1_z0 : (ehrnat_z0 [I] ulnat_z0) @ 1 =
1 @ urnat_z0),
(fun (q : 1 @ 1 = 1 @ y1)
(s : (ehrnat_y1 [I] ulnat_y1) @ 1 = 1 @ q) =>
(fun (q0 : 1 @ 1 = 1 @ y0)
(s0 : (ehrnat_y0 [I] ulnat_y0) @ 1 = 1 @ q0) =>
(fun (q1 : (wly0 @ wlz0) @ 1 = 1 @ (y0 @ z0))
(s1 : ulnat_y0 [-] ulnat_z0 = whiskerR 1 1 @ q1)
=>
(fun (q2 : (wly1 @ wlz1) @ 1 = 1 @ (y1 @ z1))
(s2 : ulnat_y1 [-] ulnat_z1 = whiskerR 1 1 @ q2)
=>
(fun (q3 : (1 @ 1 ) @ 1 = 1 @ (wly1 @ wlz1))
(s3 : (ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 1 =
whiskerR 1 1 @ q3) =>
(fun (q4 : (1 @ 1 ) @ 1 = 1 @ (wly0 @ wlz0))
(s4 : (ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 1 =
whiskerR 1 1 @ q4) =>
(fun (q5 : (...) @ wrx2 = wrx0 @ (...))
(s5 : (...) @ whiskerL wrx0 1 =
whiskerR 1 wrx2 @ q5) =>
(fun (q6 : ...) (s6 : ...) => (...) (...) (...))
((whiskerR 1 1 )^ @ (q [-] urnat_z1))
(equiv_moveL_Mp
(...^ @ ...)
(q [-] urnat_z1)
(whiskerR 1 1 ) 1 ))
((whiskerR 1 wrx2)^ @
((wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 1 ))
(equiv_moveL_Mp
((whiskerR 1 wrx2)^ @
((...) @ whiskerL wrx0 1 ))
((wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 1 )
(whiskerR 1 wrx2) 1 ))
((whiskerR 1 1 )^ @
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 1 ))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 1 ))
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 1 )
(whiskerR 1 1 ) 1 ))
((whiskerR 1 1 )^ @
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 1 ))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 1 ))
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 1 )
(whiskerR 1 1 ) 1 ))
((whiskerR 1 1 )^ @ (ulnat_y1 [-] ulnat_z1))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @ (ulnat_y1 [-] ulnat_z1))
(ulnat_y1 [-] ulnat_z1)
(whiskerR 1 1 ) 1 ))
((whiskerR 1 1 )^ @ (ulnat_y0 [-] ulnat_z0))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @ (ulnat_y0 [-] ulnat_z0))
(ulnat_y0 [-] ulnat_z0)
(whiskerR 1 1 ) 1 ))
(ehrnat_y0 [I] ulnat_y0)
(rlucancel 1 )) (ehrnat_y1 [I] ulnat_y1)
(rlucancel 1 )
snapply equiv_path_ind_rlucancel. X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, x0, wrx0, wlx1, x1, wrx1, wlx2, x2, wrx2 : a = b wly0, y0 : b = b wly1, y1 : a = a wlz0, z0 : b = b wlz1, z1 : a = a urnat_x0 : wrx0 @ 1 = 1 @ x0 urnat_x1 : wrx1 @ 1 = 1 @ x1 urnat_x2 : wrx2 @ 1 = 1 @ x2 ulnat_y0 : wly0 @ 1 = 1 @ y0 ulnat_y1 : wly1 @ 1 = 1 @ y1 ulnat_z0 : wlz0 @ 1 = 1 @ z0 ulnat_z1 : wlz1 @ 1 = 1 @ z1 urnat_z1 : 1 @ 1 = 1 @ z1ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : 1 @ 1 = 1 @ wly0ehrnat_y1 : 1 @ 1 = 1 @ wly1ehrnat_z0 : 1 @ 1 = 1 @ wlz0ehrnat_z1 : 1 @ 1 = 1 @ wlz1wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0 ehrnat_p1_z1 : (ehrnat_z1 [I] ulnat_z1) @ 1 =
1 @ urnat_z1
(fun (q : 1 @ 1 = 1 @ z0)
(s : (ehrnat_z0 [I] ulnat_z0) @ 1 = 1 @ q) =>
(fun (q0 : 1 @ 1 = 1 @ y1)
(s0 : (ehrnat_y1 [I] ulnat_y1) @ 1 = 1 @ q0) =>
(fun (q1 : 1 @ 1 = 1 @ y0)
(s1 : (ehrnat_y0 [I] ulnat_y0) @ 1 = 1 @ q1) =>
(fun (q2 : (wly0 @ wlz0) @ 1 = 1 @ (y0 @ z0))
(s2 : ulnat_y0 [-] ulnat_z0 = whiskerR 1 1 @ q2)
=>
(fun (q3 : (wly1 @ wlz1) @ 1 = 1 @ (y1 @ z1))
(s3 : ulnat_y1 [-] ulnat_z1 = whiskerR 1 1 @ q3)
=>
(fun (q4 : (1 @ 1 ) @ 1 = 1 @ (wly1 @ wlz1))
(s4 : (ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 1 =
whiskerR 1 1 @ q4) =>
(fun (q5 : (1 @ 1 ) @ 1 = 1 @ (wly0 @ wlz0))
(s5 : (ehrnat_y0 [-] ehrnat_z0) @
whiskerL 1 1 =
whiskerR 1 1 @ q5) =>
(fun (q6 : ... = ...) (s6 : ... = ...) =>
(... => ...) (...^ @ ...)
(equiv_moveL_Mp ... ... ... 1 ))
((whiskerR 1 wrx2)^ @
((...) @ whiskerL wrx0 1 ))
(equiv_moveL_Mp
((...)^ @ (...))
((...) @ whiskerL wrx0 1 )
(whiskerR 1 wrx2) 1 ))
((whiskerR 1 1 )^ @
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 1 ))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 1 ))
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 1 )
(whiskerR 1 1 ) 1 ))
((whiskerR 1 1 )^ @
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 1 ))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 1 ))
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 1 )
(whiskerR 1 1 ) 1 ))
((whiskerR 1 1 )^ @ (ulnat_y1 [-] ulnat_z1))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @ (ulnat_y1 [-] ulnat_z1))
(ulnat_y1 [-] ulnat_z1)
(whiskerR 1 1 ) 1 ))
((whiskerR 1 1 )^ @ (ulnat_y0 [-] ulnat_z0))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @ (ulnat_y0 [-] ulnat_z0))
(ulnat_y0 [-] ulnat_z0)
(whiskerR 1 1 ) 1 ))
(ehrnat_y0 [I] ulnat_y0)
(rlucancel 1 )) (ehrnat_y1 [I] ulnat_y1)
(rlucancel 1 )) (ehrnat_z0 [I] ulnat_z0)
(rlucancel 1 )
revert urnat_z1 ehrnat_p1_z1.X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, x0, wrx0, wlx1, x1, wrx1, wlx2, x2, wrx2 : a = b wly0, y0 : b = b wly1, y1 : a = a wlz0, z0 : b = b wlz1, z1 : a = a urnat_x0 : wrx0 @ 1 = 1 @ x0 urnat_x1 : wrx1 @ 1 = 1 @ x1 urnat_x2 : wrx2 @ 1 = 1 @ x2 ulnat_y0 : wly0 @ 1 = 1 @ y0 ulnat_y1 : wly1 @ 1 = 1 @ y1 ulnat_z0 : wlz0 @ 1 = 1 @ z0 ulnat_z1 : wlz1 @ 1 = 1 @ z1 ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : 1 @ 1 = 1 @ wly0ehrnat_y1 : 1 @ 1 = 1 @ wly1ehrnat_z0 : 1 @ 1 = 1 @ wlz0ehrnat_z1 : 1 @ 1 = 1 @ wlz1wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0
forall (urnat_z1 : 1 @ 1 = 1 @ z1)
(ehrnat_p1_z1 : (ehrnat_z1 [I] ulnat_z1) @ 1 =
1 @ urnat_z1),
(fun (q : 1 @ 1 = 1 @ z0)
(s : (ehrnat_z0 [I] ulnat_z0) @ 1 = 1 @ q) =>
(fun (q0 : 1 @ 1 = 1 @ y1)
(s0 : (ehrnat_y1 [I] ulnat_y1) @ 1 = 1 @ q0) =>
(fun (q1 : 1 @ 1 = 1 @ y0)
(s1 : (ehrnat_y0 [I] ulnat_y0) @ 1 = 1 @ q1) =>
(fun (q2 : (wly0 @ wlz0) @ 1 = 1 @ (y0 @ z0))
(s2 : ulnat_y0 [-] ulnat_z0 = whiskerR 1 1 @ q2)
=>
(fun (q3 : (wly1 @ wlz1) @ 1 = 1 @ (y1 @ z1))
(s3 : ulnat_y1 [-] ulnat_z1 = whiskerR 1 1 @ q3)
=>
(fun (q4 : (1 @ 1 ) @ 1 = 1 @ (wly1 @ wlz1))
(s4 : (ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 1 =
whiskerR 1 1 @ q4) =>
(fun (q5 : (...) @ 1 = 1 @ (...))
(s5 : (...) @ whiskerL 1 1 =
whiskerR 1 1 @ q5) =>
(fun (q6 : ...) (s6 : ...) => (...) (...) (...))
((whiskerR 1 wrx2)^ @ (... @ ...))
(equiv_moveL_Mp
(...^ @ ...)
(... @ ...) (whiskerR 1 wrx2) 1 ))
((whiskerR 1 1 )^ @
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 1 ))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @ ((...) @ whiskerL 1 1 ))
((ehrnat_y0 [-] ehrnat_z0) @ whiskerL 1 1 )
(whiskerR 1 1 ) 1 ))
((whiskerR 1 1 )^ @
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 1 ))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 1 ))
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 1 )
(whiskerR 1 1 ) 1 ))
((whiskerR 1 1 )^ @ (ulnat_y1 [-] ulnat_z1))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @ (ulnat_y1 [-] ulnat_z1))
(ulnat_y1 [-] ulnat_z1)
(whiskerR 1 1 ) 1 ))
((whiskerR 1 1 )^ @ (ulnat_y0 [-] ulnat_z0))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @ (ulnat_y0 [-] ulnat_z0))
(ulnat_y0 [-] ulnat_z0)
(whiskerR 1 1 ) 1 ))
(ehrnat_y0 [I] ulnat_y0)
(rlucancel 1 )) (ehrnat_y1 [I] ulnat_y1)
(rlucancel 1 )) (ehrnat_z0 [I] ulnat_z0)
(rlucancel 1 )
snapply equiv_path_ind_rlucancel. X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, x0, wrx0, wlx1, x1, wrx1, wlx2, x2, wrx2 : a = b wly0, y0 : b = b wly1, y1 : a = a wlz0, z0 : b = b wlz1, z1 : a = a urnat_x0 : wrx0 @ 1 = 1 @ x0 urnat_x1 : wrx1 @ 1 = 1 @ x1 urnat_x2 : wrx2 @ 1 = 1 @ x2 ulnat_y0 : wly0 @ 1 = 1 @ y0 ulnat_y1 : wly1 @ 1 = 1 @ y1 ulnat_z0 : wlz0 @ 1 = 1 @ z0 ulnat_z1 : wlz1 @ 1 = 1 @ z1 ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : 1 @ 1 = 1 @ wly0ehrnat_y1 : 1 @ 1 = 1 @ wly1ehrnat_z0 : 1 @ 1 = 1 @ wlz0ehrnat_z1 : 1 @ 1 = 1 @ wlz1wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0
(fun (q : 1 @ 1 = 1 @ z1)
(s : (ehrnat_z1 [I] ulnat_z1) @ 1 = 1 @ q) =>
(fun (q0 : 1 @ 1 = 1 @ z0)
(s0 : (ehrnat_z0 [I] ulnat_z0) @ 1 = 1 @ q0) =>
(fun (q1 : 1 @ 1 = 1 @ y1)
(s1 : (ehrnat_y1 [I] ulnat_y1) @ 1 = 1 @ q1) =>
(fun (q2 : 1 @ 1 = 1 @ y0)
(s2 : (ehrnat_y0 [I] ulnat_y0) @ 1 = 1 @ q2) =>
(fun (q3 : (wly0 @ wlz0) @ 1 = 1 @ (y0 @ z0))
(s3 : ulnat_y0 [-] ulnat_z0 = whiskerR 1 1 @ q3)
=>
(fun (q4 : (wly1 @ wlz1) @ 1 = 1 @ (y1 @ z1))
(s4 : ulnat_y1 [-] ulnat_z1 =
whiskerR 1 1 @ q4) =>
(fun (q5 : (1 @ 1 ) @ 1 = 1 @ (wly1 @ wlz1))
(s5 : (ehrnat_y1 [-] ehrnat_z1) @
whiskerL 1 1 =
whiskerR 1 1 @ q5) =>
(fun (q6 : ... = ...) (s6 : ... = ...) =>
(... => ...) (...^ @ ...)
(equiv_moveL_Mp ... ... ... 1 ))
((whiskerR 1 1 )^ @ ((...) @ whiskerL 1 1 ))
(equiv_moveL_Mp
((...)^ @ (...))
((...) @ whiskerL 1 1 )
(whiskerR 1 1 ) 1 ))
((whiskerR 1 1 )^ @
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 1 ))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 1 ))
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 1 )
(whiskerR 1 1 ) 1 ))
((whiskerR 1 1 )^ @ (ulnat_y1 [-] ulnat_z1))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @ (ulnat_y1 [-] ulnat_z1))
(ulnat_y1 [-] ulnat_z1)
(whiskerR 1 1 ) 1 ))
((whiskerR 1 1 )^ @ (ulnat_y0 [-] ulnat_z0))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @ (ulnat_y0 [-] ulnat_z0))
(ulnat_y0 [-] ulnat_z0)
(whiskerR 1 1 ) 1 ))
(ehrnat_y0 [I] ulnat_y0)
(rlucancel 1 )) (ehrnat_y1 [I] ulnat_y1)
(rlucancel 1 )) (ehrnat_z0 [I] ulnat_z0)
(rlucancel 1 )) (ehrnat_z1 [I] ulnat_z1)
(rlucancel 1 )
revert x0 urnat_x0.X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, wrx0, wlx1, x1, wrx1, wlx2, x2, wrx2 : a = b wly0, y0 : b = b wly1, y1 : a = a wlz0, z0 : b = b wlz1, z1 : a = a urnat_x1 : wrx1 @ 1 = 1 @ x1 urnat_x2 : wrx2 @ 1 = 1 @ x2 ulnat_y0 : wly0 @ 1 = 1 @ y0 ulnat_y1 : wly1 @ 1 = 1 @ y1 ulnat_z0 : wlz0 @ 1 = 1 @ z0 ulnat_z1 : wlz1 @ 1 = 1 @ z1 ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : 1 @ 1 = 1 @ wly0ehrnat_y1 : 1 @ 1 = 1 @ wly1ehrnat_z0 : 1 @ 1 = 1 @ wlz0ehrnat_z1 : 1 @ 1 = 1 @ wlz1wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0
forall (x0 : a = b) (urnat_x0 : wrx0 @ 1 = 1 @ x0),
(fun (q : 1 @ 1 = 1 @ z1)
(s : (ehrnat_z1 [I] ulnat_z1) @ 1 = 1 @ q) =>
(fun (q0 : 1 @ 1 = 1 @ z0)
(s0 : (ehrnat_z0 [I] ulnat_z0) @ 1 = 1 @ q0) =>
(fun (q1 : 1 @ 1 = 1 @ y1)
(s1 : (ehrnat_y1 [I] ulnat_y1) @ 1 = 1 @ q1) =>
(fun (q2 : 1 @ 1 = 1 @ y0)
(s2 : (ehrnat_y0 [I] ulnat_y0) @ 1 = 1 @ q2) =>
(fun (q3 : (wly0 @ wlz0) @ 1 = 1 @ (y0 @ z0))
(s3 : ulnat_y0 [-] ulnat_z0 = whiskerR 1 1 @ q3)
=>
(fun (q4 : (wly1 @ wlz1) @ 1 = 1 @ (y1 @ z1))
(s4 : ulnat_y1 [-] ulnat_z1 =
whiskerR 1 1 @ q4) =>
(fun (q5 : (...) @ 1 = 1 @ (...))
(s5 : (...) @ whiskerL 1 1 =
whiskerR 1 1 @ q5) =>
(fun (q6 : ...) (s6 : ...) => (...) (...) (...))
((whiskerR 1 1 )^ @ (... @ ...))
(equiv_moveL_Mp
(...^ @ ...)
(... @ ...) (whiskerR 1 1 ) 1 ))
((whiskerR 1 1 )^ @
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 1 ))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @ ((...) @ whiskerL 1 1 ))
((ehrnat_y1 [-] ehrnat_z1) @ whiskerL 1 1 )
(whiskerR 1 1 ) 1 ))
((whiskerR 1 1 )^ @ (ulnat_y1 [-] ulnat_z1))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @ (ulnat_y1 [-] ulnat_z1))
(ulnat_y1 [-] ulnat_z1)
(whiskerR 1 1 ) 1 ))
((whiskerR 1 1 )^ @ (ulnat_y0 [-] ulnat_z0))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @ (ulnat_y0 [-] ulnat_z0))
(ulnat_y0 [-] ulnat_z0)
(whiskerR 1 1 ) 1 ))
(ehrnat_y0 [I] ulnat_y0)
(rlucancel 1 )) (ehrnat_y1 [I] ulnat_y1)
(rlucancel 1 )) (ehrnat_z0 [I] ulnat_z0)
(rlucancel 1 )) (ehrnat_z1 [I] ulnat_z1)
(rlucancel 1 )
snapply equiv_path_ind_rlucancel. X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, wrx0, wlx1, x1, wrx1, wlx2, x2, wrx2 : a = b wly0, y0 : b = b wly1, y1 : a = a wlz0, z0 : b = b wlz1, z1 : a = a urnat_x1 : wrx1 @ 1 = 1 @ x1 urnat_x2 : wrx2 @ 1 = 1 @ x2 ulnat_y0 : wly0 @ 1 = 1 @ y0 ulnat_y1 : wly1 @ 1 = 1 @ y1 ulnat_z0 : wlz0 @ 1 = 1 @ z0 ulnat_z1 : wlz1 @ 1 = 1 @ z1 ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : 1 @ 1 = 1 @ wly0ehrnat_y1 : 1 @ 1 = 1 @ wly1ehrnat_z0 : 1 @ 1 = 1 @ wlz0ehrnat_z1 : 1 @ 1 = 1 @ wlz1wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0
(fun (q : a = b) (s : wrx0 @ 1 = 1 @ q) =>
(fun (q0 : 1 @ 1 = 1 @ z1)
(s0 : (ehrnat_z1 [I] ulnat_z1) @ 1 = 1 @ q0) =>
(fun (q1 : 1 @ 1 = 1 @ z0)
(s1 : (ehrnat_z0 [I] ulnat_z0) @ 1 = 1 @ q1) =>
(fun (q2 : 1 @ 1 = 1 @ y1)
(s2 : (ehrnat_y1 [I] ulnat_y1) @ 1 = 1 @ q2) =>
(fun (q3 : 1 @ 1 = 1 @ y0)
(s3 : (ehrnat_y0 [I] ulnat_y0) @ 1 = 1 @ q3) =>
(fun (q4 : (wly0 @ wlz0) @ 1 = 1 @ (y0 @ z0))
(s4 : ulnat_y0 [-] ulnat_z0 =
whiskerR 1 1 @ q4) =>
(fun (q5 : (wly1 @ wlz1) @ 1 = 1 @ (y1 @ z1))
(s5 : ulnat_y1 [-] ulnat_z1 =
whiskerR 1 1 @ q5) =>
(fun (q6 : ... = ...) (s6 : ... = ...) =>
(... => ...) (...^ @ ...)
(equiv_moveL_Mp ... ... ... 1 ))
((whiskerR 1 1 )^ @ ((...) @ whiskerL 1 1 ))
(equiv_moveL_Mp
((...)^ @ (...))
((...) @ whiskerL 1 1 )
(whiskerR 1 1 ) 1 ))
((whiskerR 1 1 )^ @ (ulnat_y1 [-] ulnat_z1))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @ (ulnat_y1 [-] ulnat_z1))
(ulnat_y1 [-] ulnat_z1)
(whiskerR 1 1 ) 1 ))
((whiskerR 1 1 )^ @ (ulnat_y0 [-] ulnat_z0))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @ (ulnat_y0 [-] ulnat_z0))
(ulnat_y0 [-] ulnat_z0)
(whiskerR 1 1 ) 1 ))
(ehrnat_y0 [I] ulnat_y0)
(rlucancel 1 )) (ehrnat_y1 [I] ulnat_y1)
(rlucancel 1 )) (ehrnat_z0 [I] ulnat_z0)
(rlucancel 1 )) (ehrnat_z1 [I] ulnat_z1)
(rlucancel 1 )) wrx0 (rlucancel 1 )
revert x1 urnat_x1.X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, wrx0, wlx1, wrx1, wlx2, x2, wrx2 : a = b wly0, y0 : b = b wly1, y1 : a = a wlz0, z0 : b = b wlz1, z1 : a = a urnat_x2 : wrx2 @ 1 = 1 @ x2 ulnat_y0 : wly0 @ 1 = 1 @ y0 ulnat_y1 : wly1 @ 1 = 1 @ y1 ulnat_z0 : wlz0 @ 1 = 1 @ z0 ulnat_z1 : wlz1 @ 1 = 1 @ z1 ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : 1 @ 1 = 1 @ wly0ehrnat_y1 : 1 @ 1 = 1 @ wly1ehrnat_z0 : 1 @ 1 = 1 @ wlz0ehrnat_z1 : 1 @ 1 = 1 @ wlz1wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0
forall (x1 : a = b) (urnat_x1 : wrx1 @ 1 = 1 @ x1),
(fun (q : a = b) (s : wrx0 @ 1 = 1 @ q) =>
(fun (q0 : 1 @ 1 = 1 @ z1)
(s0 : (ehrnat_z1 [I] ulnat_z1) @ 1 = 1 @ q0) =>
(fun (q1 : 1 @ 1 = 1 @ z0)
(s1 : (ehrnat_z0 [I] ulnat_z0) @ 1 = 1 @ q1) =>
(fun (q2 : 1 @ 1 = 1 @ y1)
(s2 : (ehrnat_y1 [I] ulnat_y1) @ 1 = 1 @ q2) =>
(fun (q3 : 1 @ 1 = 1 @ y0)
(s3 : (ehrnat_y0 [I] ulnat_y0) @ 1 = 1 @ q3) =>
(fun (q4 : (wly0 @ wlz0) @ 1 = 1 @ (y0 @ z0))
(s4 : ulnat_y0 [-] ulnat_z0 =
whiskerR 1 1 @ q4) =>
(fun (q5 : (...) @ 1 = 1 @ (...))
(s5 : ulnat_y1 [-] ulnat_z1 =
whiskerR 1 1 @ q5) =>
(fun (q6 : ...) (s6 : ...) => (...) (...) (...))
((whiskerR 1 1 )^ @ (... @ ...))
(equiv_moveL_Mp
(...^ @ ...)
(... @ ...) (whiskerR 1 1 ) 1 ))
((whiskerR 1 1 )^ @ (ulnat_y1 [-] ulnat_z1))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @ (ulnat_y1 [-] ulnat_z1))
(ulnat_y1 [-] ulnat_z1)
(whiskerR 1 1 ) 1 ))
((whiskerR 1 1 )^ @ (ulnat_y0 [-] ulnat_z0))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @ (ulnat_y0 [-] ulnat_z0))
(ulnat_y0 [-] ulnat_z0)
(whiskerR 1 1 ) 1 ))
(ehrnat_y0 [I] ulnat_y0)
(rlucancel 1 )) (ehrnat_y1 [I] ulnat_y1)
(rlucancel 1 )) (ehrnat_z0 [I] ulnat_z0)
(rlucancel 1 )) (ehrnat_z1 [I] ulnat_z1)
(rlucancel 1 )) wrx0 (rlucancel 1 )
snapply equiv_path_ind_rlucancel. X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, wrx0, wlx1, wrx1, wlx2, x2, wrx2 : a = b wly0, y0 : b = b wly1, y1 : a = a wlz0, z0 : b = b wlz1, z1 : a = a urnat_x2 : wrx2 @ 1 = 1 @ x2 ulnat_y0 : wly0 @ 1 = 1 @ y0 ulnat_y1 : wly1 @ 1 = 1 @ y1 ulnat_z0 : wlz0 @ 1 = 1 @ z0 ulnat_z1 : wlz1 @ 1 = 1 @ z1 ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : 1 @ 1 = 1 @ wly0ehrnat_y1 : 1 @ 1 = 1 @ wly1ehrnat_z0 : 1 @ 1 = 1 @ wlz0ehrnat_z1 : 1 @ 1 = 1 @ wlz1wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0
(fun (q : a = b) (s : wrx1 @ 1 = 1 @ q) =>
(fun (q0 : a = b) (s0 : wrx0 @ 1 = 1 @ q0) =>
(fun (q1 : 1 @ 1 = 1 @ z1)
(s1 : (ehrnat_z1 [I] ulnat_z1) @ 1 = 1 @ q1) =>
(fun (q2 : 1 @ 1 = 1 @ z0)
(s2 : (ehrnat_z0 [I] ulnat_z0) @ 1 = 1 @ q2) =>
(fun (q3 : 1 @ 1 = 1 @ y1)
(s3 : (ehrnat_y1 [I] ulnat_y1) @ 1 = 1 @ q3) =>
(fun (q4 : 1 @ 1 = 1 @ y0)
(s4 : (ehrnat_y0 [I] ulnat_y0) @ 1 = 1 @ q4)
=>
(fun (q5 : (wly0 @ wlz0) @ 1 = 1 @ (y0 @ z0))
(s5 : ulnat_y0 [-] ulnat_z0 =
whiskerR 1 1 @ q5) =>
(fun (q6 : ... = ...) (s6 : ... = ...) =>
(... => ...) (...^ @ ...)
(equiv_moveL_Mp ... ... ... 1 ))
((whiskerR 1 1 )^ @ (ulnat_y1 [-] ulnat_z1))
(equiv_moveL_Mp
((...)^ @ (...))
(ulnat_y1 [-] ulnat_z1)
(whiskerR 1 1 ) 1 ))
((whiskerR 1 1 )^ @ (ulnat_y0 [-] ulnat_z0))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @ (ulnat_y0 [-] ulnat_z0))
(ulnat_y0 [-] ulnat_z0)
(whiskerR 1 1 ) 1 ))
(ehrnat_y0 [I] ulnat_y0)
(rlucancel 1 )) (ehrnat_y1 [I] ulnat_y1)
(rlucancel 1 )) (ehrnat_z0 [I] ulnat_z0)
(rlucancel 1 )) (ehrnat_z1 [I] ulnat_z1)
(rlucancel 1 )) wrx0 (rlucancel 1 )) wrx1
(rlucancel 1 )
revert x2 urnat_x2.X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, wrx0, wlx1, wrx1, wlx2, wrx2 : a = b wly0, y0 : b = b wly1, y1 : a = a wlz0, z0 : b = b wlz1, z1 : a = a ulnat_y0 : wly0 @ 1 = 1 @ y0 ulnat_y1 : wly1 @ 1 = 1 @ y1 ulnat_z0 : wlz0 @ 1 = 1 @ z0 ulnat_z1 : wlz1 @ 1 = 1 @ z1 ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : 1 @ 1 = 1 @ wly0ehrnat_y1 : 1 @ 1 = 1 @ wly1ehrnat_z0 : 1 @ 1 = 1 @ wlz0ehrnat_z1 : 1 @ 1 = 1 @ wlz1wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0
forall (x2 : a = b) (urnat_x2 : wrx2 @ 1 = 1 @ x2),
(fun (q : a = b) (s : wrx1 @ 1 = 1 @ q) =>
(fun (q0 : a = b) (s0 : wrx0 @ 1 = 1 @ q0) =>
(fun (q1 : 1 @ 1 = 1 @ z1)
(s1 : (ehrnat_z1 [I] ulnat_z1) @ 1 = 1 @ q1) =>
(fun (q2 : 1 @ 1 = 1 @ z0)
(s2 : (ehrnat_z0 [I] ulnat_z0) @ 1 = 1 @ q2) =>
(fun (q3 : 1 @ 1 = 1 @ y1)
(s3 : (ehrnat_y1 [I] ulnat_y1) @ 1 = 1 @ q3) =>
(fun (q4 : 1 @ 1 = 1 @ y0)
(s4 : (ehrnat_y0 [I] ulnat_y0) @ 1 = 1 @ q4)
=>
(fun (q5 : (...) @ 1 = 1 @ (...))
(s5 : ulnat_y0 [-] ulnat_z0 =
whiskerR 1 1 @ q5) =>
(fun (q6 : ...) (s6 : ...) => (...) (...) (...))
((whiskerR 1 1 )^ @ (ulnat_y1 [-] ulnat_z1))
(equiv_moveL_Mp
(...^ @ ...)
(ulnat_y1 [-] ulnat_z1)
(whiskerR 1 1 ) 1 ))
((whiskerR 1 1 )^ @ (ulnat_y0 [-] ulnat_z0))
(equiv_moveL_Mp
((whiskerR 1 1 )^ @ (ulnat_y0 [-] ulnat_z0))
(ulnat_y0 [-] ulnat_z0)
(whiskerR 1 1 ) 1 ))
(ehrnat_y0 [I] ulnat_y0)
(rlucancel 1 )) (ehrnat_y1 [I] ulnat_y1)
(rlucancel 1 )) (ehrnat_z0 [I] ulnat_z0)
(rlucancel 1 )) (ehrnat_z1 [I] ulnat_z1)
(rlucancel 1 )) wrx0 (rlucancel 1 )) wrx1
(rlucancel 1 )
snapply equiv_path_ind_rlucancel. X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, wrx0, wlx1, wrx1, wlx2, wrx2 : a = b wly0, y0 : b = b wly1, y1 : a = a wlz0, z0 : b = b wlz1, z1 : a = a ulnat_y0 : wly0 @ 1 = 1 @ y0 ulnat_y1 : wly1 @ 1 = 1 @ y1 ulnat_z0 : wlz0 @ 1 = 1 @ z0 ulnat_z1 : wlz1 @ 1 = 1 @ z1 ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : 1 @ 1 = 1 @ wly0ehrnat_y1 : 1 @ 1 = 1 @ wly1ehrnat_z0 : 1 @ 1 = 1 @ wlz0ehrnat_z1 : 1 @ 1 = 1 @ wlz1wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0
(fun (q : a = b) (s : wrx2 @ 1 = 1 @ q) =>
(fun (q0 : a = b) (s0 : wrx1 @ 1 = 1 @ q0) =>
(fun (q1 : a = b) (s1 : wrx0 @ 1 = 1 @ q1) =>
(fun (q2 : 1 @ 1 = 1 @ z1)
(s2 : (ehrnat_z1 [I] ulnat_z1) @ 1 = 1 @ q2) =>
(fun (q3 : 1 @ 1 = 1 @ z0)
(s3 : (ehrnat_z0 [I] ulnat_z0) @ 1 = 1 @ q3) =>
(fun (q4 : 1 @ 1 = 1 @ y1)
(s4 : (ehrnat_y1 [I] ulnat_y1) @ 1 = 1 @ q4)
=>
(fun (q5 : 1 @ 1 = 1 @ y0)
(s5 : (ehrnat_y0 [I] ulnat_y0) @ 1 = 1 @ q5)
=>
(fun (q6 : ... = ...) (s6 : ... = ...) =>
(... => ...) (...^ @ ...)
(equiv_moveL_Mp ... ... ... 1 ))
((whiskerR 1 1 )^ @ (ulnat_y0 [-] ulnat_z0))
(equiv_moveL_Mp
((...)^ @ (...))
(ulnat_y0 [-] ulnat_z0)
(whiskerR 1 1 ) 1 ))
(ehrnat_y0 [I] ulnat_y0)
(rlucancel 1 )) (ehrnat_y1 [I] ulnat_y1)
(rlucancel 1 )) (ehrnat_z0 [I] ulnat_z0)
(rlucancel 1 )) (ehrnat_z1 [I] ulnat_z1)
(rlucancel 1 )) wrx0
(rlucancel 1 )) wrx1 (rlucancel 1 )) wrx2
(rlucancel 1 )
revert y0 ulnat_y0.X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, wrx0, wlx1, wrx1, wlx2, wrx2 : a = b wly0 : b = b wly1, y1 : a = a wlz0, z0 : b = b wlz1, z1 : a = a ulnat_y1 : wly1 @ 1 = 1 @ y1 ulnat_z0 : wlz0 @ 1 = 1 @ z0 ulnat_z1 : wlz1 @ 1 = 1 @ z1 ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : 1 @ 1 = 1 @ wly0ehrnat_y1 : 1 @ 1 = 1 @ wly1ehrnat_z0 : 1 @ 1 = 1 @ wlz0ehrnat_z1 : 1 @ 1 = 1 @ wlz1wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0
forall (y0 : b = b) (ulnat_y0 : wly0 @ 1 = 1 @ y0),
(fun (q : a = b) (s : wrx2 @ 1 = 1 @ q) =>
(fun (q0 : a = b) (s0 : wrx1 @ 1 = 1 @ q0) =>
(fun (q1 : a = b) (s1 : wrx0 @ 1 = 1 @ q1) =>
(fun (q2 : 1 @ 1 = 1 @ z1)
(s2 : (ehrnat_z1 [I] ulnat_z1) @ 1 = 1 @ q2) =>
(fun (q3 : 1 @ 1 = 1 @ z0)
(s3 : (ehrnat_z0 [I] ulnat_z0) @ 1 = 1 @ q3) =>
(fun (q4 : 1 @ 1 = 1 @ y1)
(s4 : (ehrnat_y1 [I] ulnat_y1) @ 1 = 1 @ q4)
=>
(fun (q5 : 1 @ 1 = 1 @ y0)
(s5 : (...) @ 1 = 1 @ q5) =>
(fun (q6 : ...) (s6 : ...) => (...) (...) (...))
((whiskerR 1 1 )^ @ (ulnat_y0 [-] ulnat_z0))
(equiv_moveL_Mp
(...^ @ ...)
(ulnat_y0 [-] ulnat_z0)
(whiskerR 1 1 ) 1 ))
(ehrnat_y0 [I] ulnat_y0)
(rlucancel 1 )) (ehrnat_y1 [I] ulnat_y1)
(rlucancel 1 )) (ehrnat_z0 [I] ulnat_z0)
(rlucancel 1 )) (ehrnat_z1 [I] ulnat_z1)
(rlucancel 1 )) wrx0
(rlucancel 1 )) wrx1 (rlucancel 1 )) wrx2
(rlucancel 1 )
snapply equiv_path_ind_rlucancel. X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, wrx0, wlx1, wrx1, wlx2, wrx2 : a = b wly0 : b = b wly1, y1 : a = a wlz0, z0 : b = b wlz1, z1 : a = a ulnat_y1 : wly1 @ 1 = 1 @ y1 ulnat_z0 : wlz0 @ 1 = 1 @ z0 ulnat_z1 : wlz1 @ 1 = 1 @ z1 ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : 1 @ 1 = 1 @ wly0ehrnat_y1 : 1 @ 1 = 1 @ wly1ehrnat_z0 : 1 @ 1 = 1 @ wlz0ehrnat_z1 : 1 @ 1 = 1 @ wlz1wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0
(fun (q : b = b) (s : wly0 @ 1 = 1 @ q) =>
(fun (q0 : a = b) (s0 : wrx2 @ 1 = 1 @ q0) =>
(fun (q1 : a = b) (s1 : wrx1 @ 1 = 1 @ q1) =>
(fun (q2 : a = b) (s2 : wrx0 @ 1 = 1 @ q2) =>
(fun (q3 : 1 @ 1 = 1 @ z1)
(s3 : (ehrnat_z1 [I] ulnat_z1) @ 1 = 1 @ q3) =>
(fun (q4 : 1 @ 1 = 1 @ z0)
(s4 : (ehrnat_z0 [I] ulnat_z0) @ 1 = 1 @ q4)
=>
(fun (q5 : 1 @ 1 = 1 @ y1)
(s5 : (ehrnat_y1 [I] ulnat_y1) @ 1 = 1 @ q5)
=>
(fun (q6 : ... = ...) (s6 : ... = ...) =>
(... => ...) (...^ @ ...)
(equiv_moveL_Mp ... ... ... 1 ))
(ehrnat_y0 [I] s)
(rlucancel 1 )) (ehrnat_y1 [I] ulnat_y1)
(rlucancel 1 )) (ehrnat_z0 [I] ulnat_z0)
(rlucancel 1 )) (ehrnat_z1 [I] ulnat_z1)
(rlucancel 1 )) wrx0
(rlucancel 1 )) wrx1
(rlucancel 1 )) wrx2 (rlucancel 1 )) wly0
(rlucancel 1 )
revert y1 ulnat_y1.X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, wrx0, wlx1, wrx1, wlx2, wrx2 : a = b wly0 : b = b wly1 : a = a wlz0, z0 : b = b wlz1, z1 : a = a ulnat_z0 : wlz0 @ 1 = 1 @ z0 ulnat_z1 : wlz1 @ 1 = 1 @ z1 ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : 1 @ 1 = 1 @ wly0ehrnat_y1 : 1 @ 1 = 1 @ wly1ehrnat_z0 : 1 @ 1 = 1 @ wlz0ehrnat_z1 : 1 @ 1 = 1 @ wlz1wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0
forall (y1 : a = a) (ulnat_y1 : wly1 @ 1 = 1 @ y1),
(fun (q : b = b) (s : wly0 @ 1 = 1 @ q) =>
(fun (q0 : a = b) (s0 : wrx2 @ 1 = 1 @ q0) =>
(fun (q1 : a = b) (s1 : wrx1 @ 1 = 1 @ q1) =>
(fun (q2 : a = b) (s2 : wrx0 @ 1 = 1 @ q2) =>
(fun (q3 : 1 @ 1 = 1 @ z1)
(s3 : (ehrnat_z1 [I] ulnat_z1) @ 1 = 1 @ q3) =>
(fun (q4 : 1 @ 1 = 1 @ z0)
(s4 : (ehrnat_z0 [I] ulnat_z0) @ 1 = 1 @ q4)
=>
(fun (q5 : 1 @ 1 = 1 @ y1)
(s5 : (...) @ 1 = 1 @ q5) =>
(fun (q6 : ...) (s6 : ...) => (...) (...) (...))
(ehrnat_y0 [I] s)
(rlucancel 1 )) (ehrnat_y1 [I] ulnat_y1)
(rlucancel 1 )) (ehrnat_z0 [I] ulnat_z0)
(rlucancel 1 )) (ehrnat_z1 [I] ulnat_z1)
(rlucancel 1 )) wrx0
(rlucancel 1 )) wrx1
(rlucancel 1 )) wrx2 (rlucancel 1 )) wly0
(rlucancel 1 )
snapply equiv_path_ind_rlucancel. X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, wrx0, wlx1, wrx1, wlx2, wrx2 : a = b wly0 : b = b wly1 : a = a wlz0, z0 : b = b wlz1, z1 : a = a ulnat_z0 : wlz0 @ 1 = 1 @ z0 ulnat_z1 : wlz1 @ 1 = 1 @ z1 ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : 1 @ 1 = 1 @ wly0ehrnat_y1 : 1 @ 1 = 1 @ wly1ehrnat_z0 : 1 @ 1 = 1 @ wlz0ehrnat_z1 : 1 @ 1 = 1 @ wlz1wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0
(fun (q : a = a) (s : wly1 @ 1 = 1 @ q) =>
(fun (q0 : b = b) (s0 : wly0 @ 1 = 1 @ q0) =>
(fun (q1 : a = b) (s1 : wrx2 @ 1 = 1 @ q1) =>
(fun (q2 : a = b) (s2 : wrx1 @ 1 = 1 @ q2) =>
(fun (q3 : a = b) (s3 : wrx0 @ 1 = 1 @ q3) =>
(fun (q4 : 1 @ 1 = 1 @ z1)
(s4 : (ehrnat_z1 [I] ulnat_z1) @ 1 = 1 @ q4)
=>
(fun (q5 : 1 @ 1 = 1 @ z0)
(s5 : (ehrnat_z0 [I] ulnat_z0) @ 1 = 1 @ q5)
=>
(fun (q6 : ... = ...) (s6 : ... = ...) =>
(... => ...) (ehrnat_y0 [I] s0) (rlucancel 1 ))
(ehrnat_y1 [I] s)
(rlucancel 1 )) (ehrnat_z0 [I] ulnat_z0)
(rlucancel 1 )) (ehrnat_z1 [I] ulnat_z1)
(rlucancel 1 )) wrx0
(rlucancel 1 )) wrx1
(rlucancel 1 )) wrx2
(rlucancel 1 )) wly0 (rlucancel 1 )) wly1
(rlucancel 1 )
revert z0 ulnat_z0.X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, wrx0, wlx1, wrx1, wlx2, wrx2 : a = b wly0 : b = b wly1 : a = a wlz0 : b = b wlz1, z1 : a = a ulnat_z1 : wlz1 @ 1 = 1 @ z1 ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : 1 @ 1 = 1 @ wly0ehrnat_y1 : 1 @ 1 = 1 @ wly1ehrnat_z0 : 1 @ 1 = 1 @ wlz0ehrnat_z1 : 1 @ 1 = 1 @ wlz1wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0
forall (z0 : b = b) (ulnat_z0 : wlz0 @ 1 = 1 @ z0),
(fun (q : a = a) (s : wly1 @ 1 = 1 @ q) =>
(fun (q0 : b = b) (s0 : wly0 @ 1 = 1 @ q0) =>
(fun (q1 : a = b) (s1 : wrx2 @ 1 = 1 @ q1) =>
(fun (q2 : a = b) (s2 : wrx1 @ 1 = 1 @ q2) =>
(fun (q3 : a = b) (s3 : wrx0 @ 1 = 1 @ q3) =>
(fun (q4 : 1 @ 1 = 1 @ z1)
(s4 : (ehrnat_z1 [I] ulnat_z1) @ 1 = 1 @ q4)
=>
(fun (q5 : 1 @ 1 = 1 @ z0)
(s5 : (...) @ 1 = 1 @ q5) =>
(fun (q6 : ...) (s6 : ...) => (...) (...) (...))
(ehrnat_y1 [I] s)
(rlucancel 1 )) (ehrnat_z0 [I] ulnat_z0)
(rlucancel 1 )) (ehrnat_z1 [I] ulnat_z1)
(rlucancel 1 )) wrx0
(rlucancel 1 )) wrx1
(rlucancel 1 )) wrx2
(rlucancel 1 )) wly0 (rlucancel 1 )) wly1
(rlucancel 1 )
snapply equiv_path_ind_rlucancel. X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, wrx0, wlx1, wrx1, wlx2, wrx2 : a = b wly0 : b = b wly1 : a = a wlz0 : b = b wlz1, z1 : a = a ulnat_z1 : wlz1 @ 1 = 1 @ z1 ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : 1 @ 1 = 1 @ wly0ehrnat_y1 : 1 @ 1 = 1 @ wly1ehrnat_z0 : 1 @ 1 = 1 @ wlz0ehrnat_z1 : 1 @ 1 = 1 @ wlz1wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0
(fun (q : b = b) (s : wlz0 @ 1 = 1 @ q) =>
(fun (q0 : a = a) (s0 : wly1 @ 1 = 1 @ q0) =>
(fun (q1 : b = b) (s1 : wly0 @ 1 = 1 @ q1) =>
(fun (q2 : a = b) (s2 : wrx2 @ 1 = 1 @ q2) =>
(fun (q3 : a = b) (s3 : wrx1 @ 1 = 1 @ q3) =>
(fun (q4 : a = b) (s4 : wrx0 @ 1 = 1 @ q4) =>
(fun (q5 : 1 @ 1 = 1 @ z1)
(s5 : (ehrnat_z1 [I] ulnat_z1) @ 1 = 1 @ q5)
=>
(fun (q6 : ... = ...) (s6 : ... = ...) =>
(... => ...) (ehrnat_y1 [I] s0) (rlucancel 1 ))
(ehrnat_z0 [I] s)
(rlucancel 1 )) (ehrnat_z1 [I] ulnat_z1)
(rlucancel 1 )) wrx0
(rlucancel 1 )) wrx1
(rlucancel 1 )) wrx2
(rlucancel 1 )) wly0
(rlucancel 1 )) wly1 (rlucancel 1 )) wlz0
(rlucancel 1 )
revert z1 ulnat_z1.X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, wrx0, wlx1, wrx1, wlx2, wrx2 : a = b wly0 : b = b wly1 : a = a wlz0 : b = b wlz1 : a = a ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : 1 @ 1 = 1 @ wly0ehrnat_y1 : 1 @ 1 = 1 @ wly1ehrnat_z0 : 1 @ 1 = 1 @ wlz0ehrnat_z1 : 1 @ 1 = 1 @ wlz1wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0
forall (z1 : a = a) (ulnat_z1 : wlz1 @ 1 = 1 @ z1),
(fun (q : b = b) (s : wlz0 @ 1 = 1 @ q) =>
(fun (q0 : a = a) (s0 : wly1 @ 1 = 1 @ q0) =>
(fun (q1 : b = b) (s1 : wly0 @ 1 = 1 @ q1) =>
(fun (q2 : a = b) (s2 : wrx2 @ 1 = 1 @ q2) =>
(fun (q3 : a = b) (s3 : wrx1 @ 1 = 1 @ q3) =>
(fun (q4 : a = b) (s4 : wrx0 @ 1 = 1 @ q4) =>
(fun (q5 : 1 @ 1 = 1 @ z1)
(s5 : (...) @ 1 = 1 @ q5) =>
(fun (q6 : ...) (s6 : ...) => (...) (...) (...))
(ehrnat_z0 [I] s)
(rlucancel 1 )) (ehrnat_z1 [I] ulnat_z1)
(rlucancel 1 )) wrx0
(rlucancel 1 )) wrx1
(rlucancel 1 )) wrx2
(rlucancel 1 )) wly0
(rlucancel 1 )) wly1 (rlucancel 1 )) wlz0
(rlucancel 1 )
snapply equiv_path_ind_rlucancel. X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, wrx0, wlx1, wrx1, wlx2, wrx2 : a = b wly0 : b = b wly1 : a = a wlz0 : b = b wlz1 : a = a ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : 1 @ 1 = 1 @ wly0ehrnat_y1 : 1 @ 1 = 1 @ wly1ehrnat_z0 : 1 @ 1 = 1 @ wlz0ehrnat_z1 : 1 @ 1 = 1 @ wlz1wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0 wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0
(fun (q : a = a) (s : wlz1 @ 1 = 1 @ q) =>
(fun (q0 : b = b) (s0 : wlz0 @ 1 = 1 @ q0) =>
(fun (q1 : a = a) (s1 : wly1 @ 1 = 1 @ q1) =>
(fun (q2 : b = b) (s2 : wly0 @ 1 = 1 @ q2) =>
(fun (q3 : a = b) (s3 : wrx2 @ 1 = 1 @ q3) =>
(fun (q4 : a = b) (s4 : wrx1 @ 1 = 1 @ q4) =>
(fun (q5 : a = b) (s5 : wrx0 @ 1 = 1 @ q5) =>
(fun (q6 : ... = ...) (s6 : ... = ...) =>
(... => ...) (ehrnat_z0 [I] s0) (rlucancel 1 ))
(ehrnat_z1 [I] s)
(rlucancel 1 )) wrx0
(rlucancel 1 )) wrx1
(rlucancel 1 )) wrx2
(rlucancel 1 )) wly0
(rlucancel 1 )) wly1
(rlucancel 1 )) wlz0 (rlucancel 1 )) wlz1
(rlucancel 1 )
revert wlrnat_y_x. (* Paired with wlx0 below. *) X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, wrx0, wlx1, wrx1, wlx2, wrx2 : a = b wly0 : b = b wly1 : a = a wlz0 : b = b wlz1 : a = a ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0 ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : 1 @ 1 = 1 @ wly0ehrnat_y1 : 1 @ 1 = 1 @ wly1ehrnat_z0 : 1 @ 1 = 1 @ wlz0ehrnat_z1 : 1 @ 1 = 1 @ wlz1wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0
forall wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0,
(fun (q : a = a) (s : wlz1 @ 1 = 1 @ q) =>
(fun (q0 : b = b) (s0 : wlz0 @ 1 = 1 @ q0) =>
(fun (q1 : a = a) (s1 : wly1 @ 1 = 1 @ q1) =>
(fun (q2 : b = b) (s2 : wly0 @ 1 = 1 @ q2) =>
(fun (q3 : a = b) (s3 : wrx2 @ 1 = 1 @ q3) =>
(fun (q4 : a = b) (s4 : wrx1 @ 1 = 1 @ q4) =>
(fun (q5 : a = b) (s5 : wrx0 @ 1 = 1 @ q5) =>
(fun (q6 : ...) (s6 : ...) => (...) (...) (...))
(ehrnat_z1 [I] s)
(rlucancel 1 )) wrx0
(rlucancel 1 )) wrx1
(rlucancel 1 )) wrx2
(rlucancel 1 )) wly0
(rlucancel 1 )) wly1
(rlucancel 1 )) wlz0 (rlucancel 1 )) wlz1
(rlucancel 1 )
revert wrx0 ehlnat_x0.X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, wlx1, wrx1, wlx2, wrx2 : a = b wly0 : b = b wly1 : a = a wlz0 : b = b wlz1 : a = a ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : 1 @ 1 = 1 @ wly0ehrnat_y1 : 1 @ 1 = 1 @ wly1ehrnat_z0 : 1 @ 1 = 1 @ wlz0ehrnat_z1 : 1 @ 1 = 1 @ wlz1wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0
forall (wrx0 : a = b)
(ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0)
(wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0),
(fun (q : a = a) (s : wlz1 @ 1 = 1 @ q) =>
(fun (q0 : b = b) (s0 : wlz0 @ 1 = 1 @ q0) =>
(fun (q1 : a = a) (s1 : wly1 @ 1 = 1 @ q1) =>
(fun (q2 : b = b) (s2 : wly0 @ 1 = 1 @ q2) =>
(fun (q3 : a = b) (s3 : wrx2 @ 1 = 1 @ q3) =>
(fun (q4 : a = b) (s4 : wrx1 @ 1 = 1 @ q4) =>
(fun (q5 : a = b) (s5 : wrx0 @ 1 = 1 @ q5) =>
(fun (q6 : ...) (s6 : ...) => (...) (...) (...))
(ehrnat_z1 [I] s)
(rlucancel 1 )) wrx0
(rlucancel 1 )) wrx1
(rlucancel 1 )) wrx2
(rlucancel 1 )) wly0
(rlucancel 1 )) wly1
(rlucancel 1 )) wlz0 (rlucancel 1 )) wlz1
(rlucancel 1 )
snapply equiv_path_ind_rlucancel. X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, wlx1, wrx1, wlx2, wrx2 : a = b wly0 : b = b wly1 : a = a wlz0 : b = b wlz1 : a = a ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : 1 @ 1 = 1 @ wly0ehrnat_y1 : 1 @ 1 = 1 @ wly1ehrnat_z0 : 1 @ 1 = 1 @ wlz0ehrnat_z1 : 1 @ 1 = 1 @ wlz1wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0
(fun (q : a = b) (s : wlx0 @ 1 = 1 @ q) =>
forall wlrnat_y_x : wly1 @ wrx1 = q @ wly0,
(fun (q0 : a = a) (s0 : wlz1 @ 1 = 1 @ q0) =>
(fun (q1 : b = b) (s1 : wlz0 @ 1 = 1 @ q1) =>
(fun (q2 : a = a) (s2 : wly1 @ 1 = 1 @ q2) =>
(fun (q3 : b = b) (s3 : wly0 @ 1 = 1 @ q3) =>
(fun (q4 : a = b) (s4 : wrx2 @ 1 = 1 @ q4) =>
(fun (q5 : a = b) (s5 : wrx1 @ 1 = 1 @ q5) =>
(fun (q6 : ...) (s6 : ...) => (...) (...) (...))
q (rlucancel 1 )) wrx1
(rlucancel 1 )) wrx2
(rlucancel 1 )) wly0
(rlucancel 1 )) wly1
(rlucancel 1 )) wlz0
(rlucancel 1 )) wlz1 (rlucancel 1 )) wlx0
(rlucancel 1 )
revert wlrnat_z_x. (* Paired with wlx1 below. *) X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, wlx1, wrx1, wlx2, wrx2 : a = b wly0 : b = b wly1 : a = a wlz0 : b = b wlz1 : a = a ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1 ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : 1 @ 1 = 1 @ wly0ehrnat_y1 : 1 @ 1 = 1 @ wly1ehrnat_z0 : 1 @ 1 = 1 @ wlz0ehrnat_z1 : 1 @ 1 = 1 @ wlz1
forall wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0,
(fun (q : a = b) (s : wlx0 @ 1 = 1 @ q) =>
forall wlrnat_y_x : wly1 @ wrx1 = q @ wly0,
(fun (q0 : a = a) (s0 : wlz1 @ 1 = 1 @ q0) =>
(fun (q1 : b = b) (s1 : wlz0 @ 1 = 1 @ q1) =>
(fun (q2 : a = a) (s2 : wly1 @ 1 = 1 @ q2) =>
(fun (q3 : b = b) (s3 : wly0 @ 1 = 1 @ q3) =>
(fun (q4 : a = b) (s4 : wrx2 @ 1 = 1 @ q4) =>
(fun (q5 : a = b) (s5 : wrx1 @ 1 = 1 @ q5) =>
(fun ... ... => ... ... ...) q (rlucancel 1 ))
wrx1 (rlucancel 1 )) wrx2
(rlucancel 1 )) wly0
(rlucancel 1 )) wly1
(rlucancel 1 )) wlz0
(rlucancel 1 )) wlz1 (rlucancel 1 )) wlx0
(rlucancel 1 )
revert wrx1 ehlnat_x1.X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, wlx1, wlx2, wrx2 : a = b wly0 : b = b wly1 : a = a wlz0 : b = b wlz1 : a = a ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : 1 @ 1 = 1 @ wly0ehrnat_y1 : 1 @ 1 = 1 @ wly1ehrnat_z0 : 1 @ 1 = 1 @ wlz0ehrnat_z1 : 1 @ 1 = 1 @ wlz1
forall (wrx1 : a = b)
(ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1)
(wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0),
(fun (q : a = b) (s : wlx0 @ 1 = 1 @ q) =>
forall wlrnat_y_x : wly1 @ wrx1 = q @ wly0,
(fun (q0 : a = a) (s0 : wlz1 @ 1 = 1 @ q0) =>
(fun (q1 : b = b) (s1 : wlz0 @ 1 = 1 @ q1) =>
(fun (q2 : a = a) (s2 : wly1 @ 1 = 1 @ q2) =>
(fun (q3 : b = b) (s3 : wly0 @ 1 = 1 @ q3) =>
(fun (q4 : a = b) (s4 : wrx2 @ 1 = 1 @ q4) =>
(fun (q5 : a = b) (s5 : wrx1 @ 1 = 1 @ q5) =>
(fun ... ... => ... ... ...) q (rlucancel 1 ))
wrx1 (rlucancel 1 )) wrx2
(rlucancel 1 )) wly0
(rlucancel 1 )) wly1
(rlucancel 1 )) wlz0
(rlucancel 1 )) wlz1 (rlucancel 1 )) wlx0
(rlucancel 1 )
snapply equiv_path_ind_rlucancel. X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, wlx1, wlx2, wrx2 : a = b wly0 : b = b wly1 : a = a wlz0 : b = b wlz1 : a = a ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2 ehrnat_y0 : 1 @ 1 = 1 @ wly0ehrnat_y1 : 1 @ 1 = 1 @ wly1ehrnat_z0 : 1 @ 1 = 1 @ wlz0ehrnat_z1 : 1 @ 1 = 1 @ wlz1
(fun (q : a = b) (s : wlx1 @ 1 = 1 @ q) =>
forall wlrnat_z_x : wlz1 @ wrx2 = q @ wlz0,
(fun (q0 : a = b) (s0 : wlx0 @ 1 = 1 @ q0) =>
forall wlrnat_y_x : wly1 @ q = q0 @ wly0,
(fun (q1 : a = a) (s1 : wlz1 @ 1 = 1 @ q1) =>
(fun (q2 : b = b) (s2 : wlz0 @ 1 = 1 @ q2) =>
(fun (q3 : a = a) (s3 : wly1 @ 1 = 1 @ q3) =>
(fun (q4 : b = b) (s4 : wly0 @ 1 = 1 @ q4) =>
(fun (q5 : a = b) (s5 : wrx2 @ 1 = 1 @ q5) =>
(fun ... ... => ... q0 ...) q (rlucancel 1 ))
wrx2 (rlucancel 1 )) wly0
(rlucancel 1 )) wly1
(rlucancel 1 )) wlz0
(rlucancel 1 )) wlz1
(rlucancel 1 )) wlx0 (rlucancel 1 )) wlx1
(rlucancel 1 )
revert wrx2 ehlnat_x2.X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, wlx1, wlx2 : a = b wly0 : b = b wly1 : a = a wlz0 : b = b wlz1 : a = a ehrnat_y0 : 1 @ 1 = 1 @ wly0ehrnat_y1 : 1 @ 1 = 1 @ wly1ehrnat_z0 : 1 @ 1 = 1 @ wlz0ehrnat_z1 : 1 @ 1 = 1 @ wlz1
forall (wrx2 : a = b)
(ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2),
(fun (q : a = b) (s : wlx1 @ 1 = 1 @ q) =>
forall wlrnat_z_x : wlz1 @ wrx2 = q @ wlz0,
(fun (q0 : a = b) (s0 : wlx0 @ 1 = 1 @ q0) =>
forall wlrnat_y_x : wly1 @ q = q0 @ wly0,
(fun (q1 : a = a) (s1 : wlz1 @ 1 = 1 @ q1) =>
(fun (q2 : b = b) (s2 : wlz0 @ 1 = 1 @ q2) =>
(fun (q3 : a = a) (s3 : wly1 @ 1 = 1 @ q3) =>
(fun (q4 : b = b) (s4 : wly0 @ 1 = 1 @ q4) =>
(fun (q5 : a = b) (s5 : ... = ...) =>
(... => ...) q (rlucancel 1 )) wrx2
(rlucancel 1 )) wly0
(rlucancel 1 )) wly1
(rlucancel 1 )) wlz0
(rlucancel 1 )) wlz1
(rlucancel 1 )) wlx0 (rlucancel 1 )) wlx1
(rlucancel 1 )
snapply equiv_path_ind_rlucancel. X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, wlx1, wlx2 : a = b wly0 : b = b wly1 : a = a wlz0 : b = b wlz1 : a = a ehrnat_y0 : 1 @ 1 = 1 @ wly0ehrnat_y1 : 1 @ 1 = 1 @ wly1ehrnat_z0 : 1 @ 1 = 1 @ wlz0ehrnat_z1 : 1 @ 1 = 1 @ wlz1
(fun (q : a = b) (s : wlx2 @ 1 = 1 @ q) =>
(fun (q0 : a = b) (s0 : wlx1 @ 1 = 1 @ q0) =>
forall wlrnat_z_x : wlz1 @ q = q0 @ wlz0,
(fun (q1 : a = b) (s1 : wlx0 @ 1 = 1 @ q1) =>
forall wlrnat_y_x : wly1 @ q0 = q1 @ wly0,
(fun (q2 : a = a) (s2 : wlz1 @ 1 = 1 @ q2) =>
(fun (q3 : b = b) (s3 : wlz0 @ 1 = 1 @ q3) =>
(fun (q4 : a = a) (s4 : wly1 @ 1 = 1 @ q4) =>
(fun (q5 : b = b) (s5 : wly0 @ 1 = 1 @ q5) =>
(fun ... ... => ... q0 ...) q (rlucancel 1 ))
wly0 (rlucancel 1 )) wly1
(rlucancel 1 )) wlz0
(rlucancel 1 )) wlz1
(rlucancel 1 )) wlx0
(rlucancel 1 )) wlx1 (rlucancel 1 )) wlx2
(rlucancel 1 )
revert wly0 ehrnat_y0.X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, wlx1, wlx2 : a = b wly1 : a = a wlz0 : b = b wlz1 : a = a ehrnat_y1 : 1 @ 1 = 1 @ wly1ehrnat_z0 : 1 @ 1 = 1 @ wlz0ehrnat_z1 : 1 @ 1 = 1 @ wlz1
forall (wly0 : b = b) (ehrnat_y0 : 1 @ 1 = 1 @ wly0),
(fun (q : a = b) (s : wlx2 @ 1 = 1 @ q) =>
(fun (q0 : a = b) (s0 : wlx1 @ 1 = 1 @ q0) =>
forall wlrnat_z_x : wlz1 @ q = q0 @ wlz0,
(fun (q1 : a = b) (s1 : wlx0 @ 1 = 1 @ q1) =>
forall wlrnat_y_x : wly1 @ q0 = q1 @ wly0,
(fun (q2 : a = a) (s2 : wlz1 @ 1 = 1 @ q2) =>
(fun (q3 : b = b) (s3 : wlz0 @ 1 = 1 @ q3) =>
(fun (q4 : a = a) (s4 : wly1 @ 1 = 1 @ q4) =>
(fun (q5 : b = b) (s5 : ... = ...) =>
(... => ...) q (rlucancel 1 )) wly0
(rlucancel 1 )) wly1
(rlucancel 1 )) wlz0
(rlucancel 1 )) wlz1
(rlucancel 1 )) wlx0
(rlucancel 1 )) wlx1 (rlucancel 1 )) wlx2
(rlucancel 1 )
snapply equiv_path_ind_rlucancel. X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, wlx1, wlx2 : a = b wly1 : a = a wlz0 : b = b wlz1 : a = a ehrnat_y1 : 1 @ 1 = 1 @ wly1ehrnat_z0 : 1 @ 1 = 1 @ wlz0ehrnat_z1 : 1 @ 1 = 1 @ wlz1
(fun (q : b = b) (s : 1 @ 1 = 1 @ q) =>
(fun (q0 : a = b) (s0 : wlx2 @ 1 = 1 @ q0) =>
(fun (q1 : a = b) (s1 : wlx1 @ 1 = 1 @ q1) =>
forall wlrnat_z_x : wlz1 @ q0 = q1 @ wlz0,
(fun (q2 : a = b) (s2 : wlx0 @ 1 = 1 @ q2) =>
forall wlrnat_y_x : wly1 @ q1 = q2 @ q,
(fun (q3 : a = a) (s3 : wlz1 @ 1 = 1 @ q3) =>
(fun (q4 : b = b) (s4 : wlz0 @ 1 = 1 @ q4) =>
(fun (q5 : a = a) (s5 : wly1 @ 1 = 1 @ q5) =>
(fun ... ... => ... q0 ...) q (rlucancel 1 ))
wly1 (rlucancel 1 )) wlz0
(rlucancel 1 )) wlz1
(rlucancel 1 )) wlx0
(rlucancel 1 )) wlx1
(rlucancel 1 )) wlx2 (rlucancel 1 )) 1
(rlucancel 1 )
revert wly1 ehrnat_y1.X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, wlx1, wlx2 : a = b wlz0 : b = b wlz1 : a = a ehrnat_z0 : 1 @ 1 = 1 @ wlz0ehrnat_z1 : 1 @ 1 = 1 @ wlz1
forall (wly1 : a = a) (ehrnat_y1 : 1 @ 1 = 1 @ wly1),
(fun (q : b = b) (s : 1 @ 1 = 1 @ q) =>
(fun (q0 : a = b) (s0 : wlx2 @ 1 = 1 @ q0) =>
(fun (q1 : a = b) (s1 : wlx1 @ 1 = 1 @ q1) =>
forall wlrnat_z_x : wlz1 @ q0 = q1 @ wlz0,
(fun (q2 : a = b) (s2 : wlx0 @ 1 = 1 @ q2) =>
forall wlrnat_y_x : wly1 @ q1 = q2 @ q,
(fun (q3 : a = a) (s3 : wlz1 @ 1 = 1 @ q3) =>
(fun (q4 : b = b) (s4 : wlz0 @ 1 = 1 @ q4) =>
(fun (q5 : a = a) (s5 : ... = ...) =>
(... => ...) q (rlucancel 1 )) wly1
(rlucancel 1 )) wlz0
(rlucancel 1 )) wlz1
(rlucancel 1 )) wlx0
(rlucancel 1 )) wlx1
(rlucancel 1 )) wlx2 (rlucancel 1 )) 1
(rlucancel 1 )
snapply equiv_path_ind_rlucancel. X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, wlx1, wlx2 : a = b wlz0 : b = b wlz1 : a = a ehrnat_z0 : 1 @ 1 = 1 @ wlz0ehrnat_z1 : 1 @ 1 = 1 @ wlz1
(fun (q : a = a) (s : 1 @ 1 = 1 @ q) =>
(fun (q0 : b = b) (s0 : 1 @ 1 = 1 @ q0) =>
(fun (q1 : a = b) (s1 : wlx2 @ 1 = 1 @ q1) =>
(fun (q2 : a = b) (s2 : wlx1 @ 1 = 1 @ q2) =>
forall wlrnat_z_x : wlz1 @ q1 = q2 @ wlz0,
(fun (q3 : a = b) (s3 : wlx0 @ 1 = 1 @ q3) =>
forall wlrnat_y_x : q @ q2 = q3 @ q0,
(fun (q4 : a = a) (s4 : wlz1 @ 1 = 1 @ q4) =>
(fun (q5 : b = b) (s5 : wlz0 @ 1 = 1 @ q5) =>
(fun ... ... => ... q0 ...) q (rlucancel 1 ))
wlz0 (rlucancel 1 )) wlz1
(rlucancel 1 )) wlx0
(rlucancel 1 )) wlx1
(rlucancel 1 )) wlx2
(rlucancel 1 )) 1 (rlucancel 1 )) 1
(rlucancel 1 )
revert wlz0 ehrnat_z0.X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, wlx1, wlx2 : a = b wlz1 : a = a ehrnat_z1 : 1 @ 1 = 1 @ wlz1
forall (wlz0 : b = b) (ehrnat_z0 : 1 @ 1 = 1 @ wlz0),
(fun (q : a = a) (s : 1 @ 1 = 1 @ q) =>
(fun (q0 : b = b) (s0 : 1 @ 1 = 1 @ q0) =>
(fun (q1 : a = b) (s1 : wlx2 @ 1 = 1 @ q1) =>
(fun (q2 : a = b) (s2 : wlx1 @ 1 = 1 @ q2) =>
forall wlrnat_z_x : wlz1 @ q1 = q2 @ wlz0,
(fun (q3 : a = b) (s3 : wlx0 @ 1 = 1 @ q3) =>
forall wlrnat_y_x : q @ q2 = q3 @ q0,
(fun (q4 : a = a) (s4 : wlz1 @ 1 = 1 @ q4) =>
(fun (q5 : b = b) (s5 : ... = ...) =>
(... => ...) q (rlucancel 1 )) wlz0
(rlucancel 1 )) wlz1
(rlucancel 1 )) wlx0
(rlucancel 1 )) wlx1
(rlucancel 1 )) wlx2
(rlucancel 1 )) 1 (rlucancel 1 )) 1
(rlucancel 1 )
snapply equiv_path_ind_rlucancel. X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, wlx1, wlx2 : a = b wlz1 : a = a ehrnat_z1 : 1 @ 1 = 1 @ wlz1
(fun (q : b = b) (s : 1 @ 1 = 1 @ q) =>
(fun (q0 : a = a) (s0 : 1 @ 1 = 1 @ q0) =>
(fun (q1 : b = b) (s1 : 1 @ 1 = 1 @ q1) =>
(fun (q2 : a = b) (s2 : wlx2 @ 1 = 1 @ q2) =>
(fun (q3 : a = b) (s3 : wlx1 @ 1 = 1 @ q3) =>
forall wlrnat_z_x : wlz1 @ q2 = q3 @ q,
(fun (q4 : a = b) (s4 : wlx0 @ 1 = 1 @ q4) =>
forall wlrnat_y_x : q0 @ q3 = q4 @ q1,
(fun (q5 : a = a) (s5 : wlz1 @ 1 = 1 @ q5) =>
(fun ... ... => ... q0 ...) q (rlucancel 1 ))
wlz1 (rlucancel 1 )) wlx0
(rlucancel 1 )) wlx1
(rlucancel 1 )) wlx2
(rlucancel 1 )) 1 (rlucancel 1 )) 1
(rlucancel 1 )) 1 (rlucancel 1 )
revert wlz1 ehrnat_z1.X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, wlx1, wlx2 : a = b
forall (wlz1 : a = a) (ehrnat_z1 : 1 @ 1 = 1 @ wlz1),
(fun (q : b = b) (s : 1 @ 1 = 1 @ q) =>
(fun (q0 : a = a) (s0 : 1 @ 1 = 1 @ q0) =>
(fun (q1 : b = b) (s1 : 1 @ 1 = 1 @ q1) =>
(fun (q2 : a = b) (s2 : wlx2 @ 1 = 1 @ q2) =>
(fun (q3 : a = b) (s3 : wlx1 @ 1 = 1 @ q3) =>
forall wlrnat_z_x : wlz1 @ q2 = q3 @ q,
(fun (q4 : a = b) (s4 : wlx0 @ 1 = 1 @ q4) =>
forall wlrnat_y_x : q0 @ q3 = q4 @ q1,
(fun (q5 : a = a) (s5 : ... = ...) =>
(... => ...) q (rlucancel 1 )) wlz1
(rlucancel 1 )) wlx0
(rlucancel 1 )) wlx1
(rlucancel 1 )) wlx2
(rlucancel 1 )) 1 (rlucancel 1 )) 1
(rlucancel 1 )) 1 (rlucancel 1 )
snapply equiv_path_ind_rlucancel. X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, wlx1, wlx2 : a = b
(fun (q : a = a) (s : 1 @ 1 = 1 @ q) =>
(fun (q0 : b = b) (s0 : 1 @ 1 = 1 @ q0) =>
(fun (q1 : a = a) (s1 : 1 @ 1 = 1 @ q1) =>
(fun (q2 : b = b) (s2 : 1 @ 1 = 1 @ q2) =>
(fun (q3 : a = b) (s3 : wlx2 @ 1 = 1 @ q3) =>
(fun (q4 : a = b) (s4 : wlx1 @ 1 = 1 @ q4) =>
forall wlrnat_z_x : q @ q3 = q4 @ q0,
(fun (q5 : a = b) (s5 : wlx0 @ 1 = 1 @ q5) =>
forall wlrnat_y_x : q1 @ q4 = q5 @ q2,
(fun ... ... => ... q0 ...) q (rlucancel 1 ))
wlx0 (rlucancel 1 )) wlx1
(rlucancel 1 )) wlx2
(rlucancel 1 )) 1 (rlucancel 1 )) 1
(rlucancel 1 )) 1 (rlucancel 1 )) 1
(rlucancel 1 )
revert wlx1.X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, wlx2 : a = b
forall wlx1 : a = b,
(fun (q : a = a) (s : 1 @ 1 = 1 @ q) =>
(fun (q0 : b = b) (s0 : 1 @ 1 = 1 @ q0) =>
(fun (q1 : a = a) (s1 : 1 @ 1 = 1 @ q1) =>
(fun (q2 : b = b) (s2 : 1 @ 1 = 1 @ q2) =>
(fun (q3 : a = b) (s3 : wlx2 @ 1 = 1 @ q3) =>
(fun (q4 : a = b) (s4 : wlx1 @ 1 = 1 @ q4) =>
forall wlrnat_z_x : q @ q3 = q4 @ q0,
(fun (q5 : a = b) (s5 : wlx0 @ 1 = 1 @ q5) =>
forall wlrnat_y_x : q1 @ q4 = q5 @ q2,
(... => ...) q (rlucancel 1 )) wlx0
(rlucancel 1 )) wlx1
(rlucancel 1 )) wlx2
(rlucancel 1 )) 1 (rlucancel 1 )) 1
(rlucancel 1 )) 1 (rlucancel 1 )) 1
(rlucancel 1 )
snapply equiv_path_ind_lrucancel. X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx0, wlx2 : a = b
(fun (q : a = b) (s : 1 @ wlx2 = q @ 1 ) =>
(fun (q0 : a = b) (s0 : wlx0 @ 1 = 1 @ q0) =>
forall wlrnat_y_x : 1 @ q = q0 @ 1 ,
(fun (q1 : a = a) (s1 : 1 @ 1 = 1 @ q1) =>
(fun (q2 : b = b) (s2 : 1 @ 1 = 1 @ q2) =>
(fun (q3 : a = a) (s3 : 1 @ 1 = 1 @ q3) =>
(fun (q4 : b = b) (s4 : 1 @ 1 = 1 @ q4) =>
(fun (q5 : a = b) (s5 : wlx2 @ 1 = 1 @ q5) =>
(fun (q6 : ...) (s6 : ...) => (...) q0 (...)) q
(rlucancel 1 )) wlx2
(rlucancel 1 )) 1
(rlucancel 1 )) 1 (rlucancel 1 )) 1
(rlucancel 1 )) 1 (rlucancel 1 )) wlx0
(rlucancel 1 )) wlx2 (lrucancel 1 )
revert wlx0.X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx2 : a = b
forall wlx0 : a = b,
(fun (q : a = b) (s : 1 @ wlx2 = q @ 1 ) =>
(fun (q0 : a = b) (s0 : wlx0 @ 1 = 1 @ q0) =>
forall wlrnat_y_x : 1 @ q = q0 @ 1 ,
(fun (q1 : a = a) (s1 : 1 @ 1 = 1 @ q1) =>
(fun (q2 : b = b) (s2 : 1 @ 1 = 1 @ q2) =>
(fun (q3 : a = a) (s3 : 1 @ 1 = 1 @ q3) =>
(fun (q4 : b = b) (s4 : 1 @ 1 = 1 @ q4) =>
(fun (q5 : a = b) (s5 : wlx2 @ 1 = 1 @ q5) =>
(fun ... ... => ... q0 ...) q (rlucancel 1 ))
wlx2 (rlucancel 1 )) 1
(rlucancel 1 )) 1 (rlucancel 1 )) 1
(rlucancel 1 )) 1 (rlucancel 1 )) wlx0
(rlucancel 1 )) wlx2 (lrucancel 1 )
snapply equiv_path_ind_lrucancel. X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal b : X wlx2 : a = b
(fun (q : a = b) (s : 1 @ wlx2 = q @ 1 ) =>
(fun (q0 : a = a) (s0 : 1 @ 1 = 1 @ q0) =>
(fun (q1 : b = b) (s1 : 1 @ 1 = 1 @ q1) =>
(fun (q2 : a = a) (s2 : 1 @ 1 = 1 @ q2) =>
(fun (q3 : b = b) (s3 : 1 @ 1 = 1 @ q3) =>
(fun (q4 : a = b) (s4 : wlx2 @ 1 = 1 @ q4) =>
(fun (q5 : a = b) (s5 : wlx2 @ 1 = 1 @ q5) =>
(fun (q6 : a = b) (s6 : ... = ...) =>
(... => ...) (... [I] s0) (rlucancel 1 )) q
(rlucancel 1 )) wlx2
(rlucancel 1 )) wlx2
(rlucancel 1 )) 1 (rlucancel 1 )) 1
(rlucancel 1 )) 1 (rlucancel 1 )) 1
(rlucancel 1 )) wlx2 (lrucancel 1 )
destruct wlx2.X : Type a : X P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3) (x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0) (x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2) (x25 x26 : x = x3)
(x27 x28 : a = x2) (x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5) (x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8) (x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6) (x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26) (x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28) (x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type H : ?Goal
P a a a a a 1 1 1 1 1 1 1 1 1 1 1 1 1 1 1 1 1 1 1 1 1
(1 @ 1 ) (1 @ 1 ) (1 @ 1 )
(1 @ 1 ) (rlucancel 1 [I] rlucancel 1 )
(rlucancel 1 ) (rlucancel 1 [I] rlucancel 1 )
(rlucancel 1 ) (rlucancel 1 [I] rlucancel 1 )
(rlucancel 1 ) (rlucancel 1 )
(rlucancel 1 [I] rlucancel 1 )
(rlucancel 1 ) (rlucancel 1 [I] rlucancel 1 )
(rlucancel 1 ) (rlucancel 1 [I] rlucancel 1 )
(rlucancel 1 ) (rlucancel 1 [I] rlucancel 1 )
((whiskerR 1 1 )^ @ (rlucancel 1 [-] rlucancel 1 ))
((whiskerR 1 1 )^ @
((rlucancel 1 [I] rlucancel 1 )
[-] (rlucancel 1 [I] rlucancel 1 )))
((whiskerR 1 1 )^ @ (rlucancel 1 [-] rlucancel 1 ))
((whiskerR 1 1 )^ @
((rlucancel 1 [I] rlucancel 1 )
[-] (rlucancel 1 [I] rlucancel 1 )))
(rlucancel 1 ) (rlucancel 1 )
(rlucancel 1 ) (rlucancel 1 )
(rlucancel 1 ) (rlucancel 1 )
(rlucancel 1 )
((whiskerR 1 1 )^ @
((rlucancel 1 [-] rlucancel 1 ) @ whiskerL 1 1 ))
((whiskerR 1 1 )^ @
((rlucancel 1 [-] rlucancel 1 ) @ whiskerL 1 1 ))
(((concat_p1 (1 @ 1 ))^ @
(((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 (lrucancel 1 )^) @
(rlucancel 1 [-] rlucancel 1 )^)) @
concat_p1 (1 @ 1 )) (lrucancel 1 )
(((concat_p1 (1 @ 1 ))^ @
(((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 (lrucancel 1 )^) @
(rlucancel 1 [-] rlucancel 1 )^)) @
concat_p1 (1 @ 1 )) (lrucancel 1 )
((whiskerL 1 1 )^ @
((((concat_p1 (1 @ 1 ))^ @
(((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 (lrucancel 1 )^) @
(rlucancel 1 [-] rlucancel 1 )^)) @
concat_p1 (1 @ 1 )
[I] ((concat_p1 (1 @ 1 ))^ @
(((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 (lrucancel 1 )^) @
(rlucancel 1 [-] rlucancel 1 )^)) @
concat_p1 (1 @ 1 )) @
whiskerR 1 1 ))
((whiskerR 1 1 )^ @
((lrucancel 1 [-] lrucancel 1 ) @ whiskerL 1 1 )) 1 1
1 1
(equiv_moveL_Mp
((whiskerR 1 1 )^ @ (rlucancel 1 [-] rlucancel 1 ))
(rlucancel 1 [-] rlucancel 1 )
(whiskerR 1 1 ) 1 )
(equiv_moveL_Mp
((whiskerR 1 1 )^ @
((rlucancel 1 [I] rlucancel 1 )
[-] (rlucancel 1 [I] rlucancel 1 )))
((rlucancel 1 [I] rlucancel 1 )
[-] (rlucancel 1 [I] rlucancel 1 ))
(whiskerR 1 1 ) 1 )
(equiv_moveL_Mp
((whiskerR 1 1 )^ @ (rlucancel 1 [-] rlucancel 1 ))
(rlucancel 1 [-] rlucancel 1 )
(whiskerR 1 1 ) 1 )
(equiv_moveL_Mp
((whiskerR 1 1 )^ @
((rlucancel 1 [I] rlucancel 1 )
[-] (rlucancel 1 [I] rlucancel 1 )))
((rlucancel 1 [I] rlucancel 1 )
[-] (rlucancel 1 [I] rlucancel 1 ))
(whiskerR 1 1 ) 1 )
(equiv_moveL_Mp
((whiskerR 1 1 )^ @
((rlucancel 1 [-] rlucancel 1 ) @ whiskerL 1 1 ))
((rlucancel 1 [-] rlucancel 1 ) @ whiskerL 1 1 )
(whiskerR 1 1 ) 1 )
(equiv_moveL_Mp
((whiskerR 1 1 )^ @
((rlucancel 1 [-] rlucancel 1 ) @ whiskerL 1 1 ))
((rlucancel 1 [-] rlucancel 1 ) @ whiskerL 1 1 )
(whiskerR 1 1 ) 1 )
(equiv_moveL_Mp
((whiskerL 1 1 )^ @
((((concat_p1 (1 @ 1 ))^ @
(((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 (lrucancel 1 )^) @
(rlucancel 1 [-] rlucancel 1 )^)) @
concat_p1 (1 @ 1 )
[I] ((concat_p1 (1 @ 1 ))^ @
(((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 (lrucancel 1 )^) @
(rlucancel 1 [-] rlucancel 1 )^)) @
concat_p1 (1 @ 1 )) @
whiskerR 1 1 ))
((((concat_p1 (1 @ 1 ))^ @
(((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 (lrucancel 1 )^) @
(rlucancel 1 [-] rlucancel 1 )^)) @
concat_p1 (1 @ 1 )
[I] ((concat_p1 (1 @ 1 ))^ @
(((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 (lrucancel 1 )^) @
(rlucancel 1 [-] rlucancel 1 )^)) @
concat_p1 (1 @ 1 )) @
whiskerR 1 1 ) (whiskerL 1 1 ) 1 )
(equiv_moveL_Mp
((whiskerR 1 1 )^ @
((lrucancel 1 [-] lrucancel 1 ) @ whiskerL 1 1 ))
((lrucancel 1 [-] lrucancel 1 ) @ whiskerL 1 1 )
(whiskerR 1 1 ) 1 ) (rlucancel 1 )
(rlucancel 1 ) (rlucancel 1 )
(rlucancel 1 ) (rlucancel 1 )
(rlucancel 1 ) (rlucancel 1 )
(equiv_helper (rlucancel 1 [-] rlucancel 1 )
((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 (lrucancel 1 )^)
(((concat_p1 (1 @ 1 ))^ @
(((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 (lrucancel 1 )^) @
(rlucancel 1 [-] rlucancel 1 )^)) @
concat_p1 (1 @ 1 )) 1 )
(equiv_helper (rlucancel 1 [-] rlucancel 1 )
((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 (lrucancel 1 )^)
(((concat_p1 (1 @ 1 ))^ @
(((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 (lrucancel 1 )^) @
(rlucancel 1 [-] rlucancel 1 )^)) @
concat_p1 (1 @ 1 )) 1 )
(* Remove the next two lines if not using the [generalize_goal] tactic. *)
exact H. } X : Type a : X lem : forall
P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3)
(x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0)
(x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2)
(x25 x26 : x = x3) (x27 x28 : a = x2)
(x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5)
(x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8)
(x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6)
(x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26)
(x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28)
(x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type ,
P a a a a a 1 1 1 1 1 1 1 1 1 1 1 1 1 1 1 1 1 1
1 1 1 (1 @ 1 ) (1 @ 1 ) (1 @ 1 )
(1 @ 1 ) (rlucancel 1 [I] rlucancel 1 )
(rlucancel 1 ) (rlucancel 1 [I] rlucancel 1 )
(rlucancel 1 ) (rlucancel 1 [I] rlucancel 1 )
(rlucancel 1 ) (rlucancel 1 )
(rlucancel 1 [I] rlucancel 1 )
(rlucancel 1 ) (rlucancel 1 [I] rlucancel 1 )
(rlucancel 1 ) (rlucancel 1 [I] rlucancel 1 )
(rlucancel 1 ) (rlucancel 1 [I] rlucancel 1 )
((whiskerR 1 1 )^ @
(rlucancel 1 [-] rlucancel 1 ))
((whiskerR 1 1 )^ @
((rlucancel 1 [I] rlucancel 1 )
[-] (rlucancel 1 [I] rlucancel 1 )))
((whiskerR 1 1 )^ @
(rlucancel 1 [-] rlucancel 1 ))
((whiskerR 1 1 )^ @
((rlucancel 1 [I] rlucancel 1 )
[-] (rlucancel 1 [I] rlucancel 1 )))
(rlucancel 1 ) (rlucancel 1 )
(rlucancel 1 ) (rlucancel 1 )
(rlucancel 1 ) (rlucancel 1 )
(rlucancel 1 )
((whiskerR 1 1 )^ @
((rlucancel 1 [-] rlucancel 1 ) @ whiskerL 1 1 ))
((whiskerR 1 1 )^ @
((rlucancel 1 [-] rlucancel 1 ) @ whiskerL 1 1 ))
(((concat_p1 (1 @ 1 ))^ @
(((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 (lrucancel 1 )^) @
(rlucancel 1 [-] rlucancel 1 )^)) @
concat_p1 (1 @ 1 )) (lrucancel 1 )
(((concat_p1 (1 @ 1 ))^ @
(((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 (lrucancel 1 )^) @
(rlucancel 1 [-] rlucancel 1 )^)) @
concat_p1 (1 @ 1 )) (lrucancel 1 )
((whiskerL 1 1 )^ @
((((concat_p1 (1 @ 1 ))^ @
(((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 (lrucancel 1 )^) @
(rlucancel 1 [-] rlucancel 1 )^)) @
concat_p1 (1 @ 1 )
[I] ((concat_p1 (1 @ 1 ))^ @
(((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 (lrucancel 1 )^) @
(rlucancel 1 [-] rlucancel 1 )^)) @
concat_p1 (1 @ 1 )) @
whiskerR 1 1 ))
((whiskerR 1 1 )^ @
((lrucancel 1 [-] lrucancel 1 ) @ whiskerL 1 1 ))
1 1 1 1
(equiv_moveL_Mp
((whiskerR 1 1 )^ @
(rlucancel 1 [-] rlucancel 1 ))
(rlucancel 1 [-] rlucancel 1 )
(whiskerR 1 1 ) 1 )
(equiv_moveL_Mp
((whiskerR 1 1 )^ @
((rlucancel 1 [I] rlucancel 1 )
[-] (rlucancel 1 [I] rlucancel 1 )))
((rlucancel 1 [I] rlucancel 1 )
[-] (rlucancel 1 [I] rlucancel 1 ))
(whiskerR 1 1 ) 1 )
(equiv_moveL_Mp
((whiskerR 1 1 )^ @
(rlucancel 1 [-] rlucancel 1 ))
(rlucancel 1 [-] rlucancel 1 )
(whiskerR 1 1 ) 1 )
(equiv_moveL_Mp
((whiskerR 1 1 )^ @
((rlucancel 1 [I] rlucancel 1 )
[-] (rlucancel 1 [I] rlucancel 1 )))
((rlucancel 1 [I] rlucancel 1 )
[-] (rlucancel 1 [I] rlucancel 1 ))
(whiskerR 1 1 ) 1 )
(equiv_moveL_Mp
((whiskerR 1 1 )^ @
((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 1 ))
((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 1 ) (whiskerR 1 1 ) 1 )
(equiv_moveL_Mp
((whiskerR 1 1 )^ @
((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 1 ))
((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 1 ) (whiskerR 1 1 ) 1 )
(equiv_moveL_Mp
((whiskerL 1 1 )^ @
((((concat_p1 (1 @ 1 ))^ @
(((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 (lrucancel 1 )^) @
(rlucancel 1 [-] rlucancel 1 )^)) @
concat_p1 (1 @ 1 )
[I] ((concat_p1 (1 @ 1 ))^ @
(((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 (lrucancel 1 )^) @
(rlucancel 1 [-] rlucancel 1 )^)) @
concat_p1 (1 @ 1 )) @
whiskerR 1 1 ))
((((concat_p1 (1 @ 1 ))^ @
(((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 (lrucancel 1 )^) @
(rlucancel 1 [-] rlucancel 1 )^)) @
concat_p1 (1 @ 1 )
[I] ((concat_p1 (1 @ 1 ))^ @
(((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 (lrucancel 1 )^) @
(rlucancel 1 [-] rlucancel 1 )^)) @
concat_p1 (1 @ 1 )) @
whiskerR 1 1 ) (whiskerL 1 1 ) 1 )
(equiv_moveL_Mp
((whiskerR 1 1 )^ @
((lrucancel 1 [-] lrucancel 1 ) @
whiskerL 1 1 ))
((lrucancel 1 [-] lrucancel 1 ) @
whiskerL 1 1 ) (whiskerR 1 1 ) 1 )
(rlucancel 1 ) (rlucancel 1 )
(rlucancel 1 ) (rlucancel 1 )
(rlucancel 1 ) (rlucancel 1 )
(rlucancel 1 )
(equiv_helper (rlucancel 1 [-] rlucancel 1 )
((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 (lrucancel 1 )^)
(((concat_p1 (1 @ 1 ))^ @
(((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 (lrucancel 1 )^) @
(rlucancel 1 [-] rlucancel 1 )^)) @
concat_p1 (1 @ 1 )) 1 )
(equiv_helper (rlucancel 1 [-] rlucancel 1 )
((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 (lrucancel 1 )^)
(((concat_p1 (1 @ 1 ))^ @
(((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 (lrucancel 1 )^) @
(rlucancel 1 [-] rlucancel 1 )^)) @
concat_p1 (1 @ 1 )) 1 ) ->
forall (b c d e f : X) (wlx0 x0 wrx0 : a = b)
(wlx1 x1 wrx1 : c = d) (wlx2 x2 wrx2 : e = f)
(wly0 y0 wry0 : b = d) (wly1 y1 wry1 : a = c)
(wlz0 z0 wrz0 : d = f) (wlz1 z1 wrz1 : c = e)
(wlyz0 wryz0 : b = f) (wlyz1 wryz1 : a = e)
(ulnat_x0 : wlx0 @ 1 = 1 @ x0)
(urnat_x0 : wrx0 @ 1 = 1 @ x0)
(ulnat_x1 : wlx1 @ 1 = 1 @ x1)
(urnat_x1 : wrx1 @ 1 = 1 @ x1)
(ulnat_x2 : wlx2 @ 1 = 1 @ x2)
(urnat_x2 : wrx2 @ 1 = 1 @ x2)
(ulnat_y0 : wly0 @ 1 = 1 @ y0)
(urnat_y0 : wry0 @ 1 = 1 @ y0)
(ulnat_y1 : wly1 @ 1 = 1 @ y1)
(urnat_y1 : wry1 @ 1 = 1 @ y1)
(ulnat_z0 : wlz0 @ 1 = 1 @ z0)
(urnat_z0 : wrz0 @ 1 = 1 @ z0)
(ulnat_z1 : wlz1 @ 1 = 1 @ z1)
(urnat_z1 : wrz1 @ 1 = 1 @ z1)
(ulnat_yz0 : wlyz0 @ 1 = 1 @ (y0 @ z0))
(urnat_yz0 : wryz0 @ 1 = 1 @ (y0 @ z0))
(ulnat_yz1 : wlyz1 @ 1 = 1 @ (y1 @ z1))
(urnat_yz1 : wryz1 @ 1 = 1 @ (y1 @ z1))
(ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0)
(ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1)
(ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2)
(ehrnat_y0 : wry0 @ 1 = 1 @ wly0)
(ehrnat_y1 : wry1 @ 1 = 1 @ wly1)
(ehrnat_z0 : wrz0 @ 1 = 1 @ wlz0)
(ehrnat_z1 : wrz1 @ 1 = 1 @ wlz1)
(ehrnat_yz0 : wryz0 @ 1 = 1 @ wlyz0)
(ehrnat_yz1 : wryz1 @ 1 = 1 @ wlyz1)
(wlrnat_x_y : wlx0 @ wry0 = wry1 @ wlx1)
(wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0)
(wlrnat_x_z : wlx1 @ wrz0 = wrz1 @ wlx2)
(wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0)
(wlrnat_x_yz : wlx0 @ wryz0 = wryz1 @ wlx2)
(wlrnat_yz_x : wlyz1 @ wrx2 = wrx0 @ wlyz0)
(wrpp_yz0 : wry0 @ wrz0 = wryz0)
(wlpp_yz0 : wly0 @ wlz0 = wlyz0)
(wrpp_yz1 : wry1 @ wrz1 = wryz1)
(wlpp_yz1 : wly1 @ wlz1 = wlyz1)
(H_ulnat_yz0 : ulnat_y0 [-] ulnat_z0 =
whiskerR wlpp_yz0 1 @ ulnat_yz0)
(H_urnat_yz0 : urnat_y0 [-] urnat_z0 =
whiskerR wrpp_yz0 1 @ urnat_yz0)
(H_ulnat_yz1 : ulnat_y1 [-] ulnat_z1 =
whiskerR wlpp_yz1 1 @ ulnat_yz1)
(H_urnat_yz1 : urnat_y1 [-] urnat_z1 =
whiskerR wrpp_yz1 1 @ urnat_yz1)
(H_ehrnat_yz0 : (ehrnat_y0 [-] ehrnat_z0) @
whiskerL 1 wlpp_yz0 =
whiskerR wrpp_yz0 1 @ ehrnat_yz0)
(H_ehrnat_yz1 : (ehrnat_y1 [-] ehrnat_z1) @
whiskerL 1 wlpp_yz1 =
whiskerR wrpp_yz1 1 @ ehrnat_yz1)
(H_wlrnat_x_yz : (wlrnat_x_y [I] wlrnat_x_z) @
whiskerR wrpp_yz1 wlx2 =
whiskerL wlx0 wrpp_yz0 @
wlrnat_x_yz)
(H_wlrnat_yz_x : (wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 wlpp_yz0 =
whiskerR wlpp_yz1 wrx2 @
wlrnat_yz_x)
(ehlnat_1p_x0 : (ehlnat_x0 [I] urnat_x0) @ 1 =
1 @ ulnat_x0)
(ehlnat_1p_x1 : (ehlnat_x1 [I] urnat_x1) @ 1 =
1 @ ulnat_x1)
(ehlnat_1p_x2 : (ehlnat_x2 [I] urnat_x2) @ 1 =
1 @ ulnat_x2)
(ehrnat_p1_y0 : (ehrnat_y0 [I] ulnat_y0) @ 1 =
1 @ urnat_y0)
(ehrnat_p1_y1 : (ehrnat_y1 [I] ulnat_y1) @ 1 =
1 @ urnat_y1)
(ehrnat_p1_z0 : (ehrnat_z0 [I] ulnat_z0) @ 1 =
1 @ urnat_z0)
(ehrnat_p1_z1 : (ehrnat_z1 [I] ulnat_z1) @ 1 =
1 @ urnat_z1)
(wlrnat_V_x_y : whiskerR wlrnat_x_y 1 @
(ehrnat_y1 [-] ehlnat_x1) =
(ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^)
(wlrnat_V_x_z : whiskerR wlrnat_x_z 1 @
(ehrnat_z1 [-] ehlnat_x2) =
(ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^),
P b c d e f wlx0 x0 wrx0 wlx1 x1 wrx1 wlx2 x2
wrx2 wly0 y0 wry0 wly1 y1 wry1 wlz0 z0 wrz0
wlz1 z1 wrz1 wlyz0 wryz0 wlyz1 wryz1 ulnat_x0
urnat_x0 ulnat_x1 urnat_x1 ulnat_x2 urnat_x2
ulnat_y0 urnat_y0 ulnat_y1 urnat_y1 ulnat_z0
urnat_z0 ulnat_z1 urnat_z1 ulnat_yz0 urnat_yz0
ulnat_yz1 urnat_yz1 ehlnat_x0 ehlnat_x1
ehlnat_x2 ehrnat_y0 ehrnat_y1 ehrnat_z0
ehrnat_z1 ehrnat_yz0 ehrnat_yz1 wlrnat_x_y
wlrnat_y_x wlrnat_x_z wlrnat_z_x wlrnat_x_yz
wlrnat_yz_x wrpp_yz0 wlpp_yz0 wrpp_yz1
wlpp_yz1 H_ulnat_yz0 H_urnat_yz0 H_ulnat_yz1
H_urnat_yz1 H_ehrnat_yz0 H_ehrnat_yz1
H_wlrnat_x_yz H_wlrnat_yz_x ehlnat_1p_x0
ehlnat_1p_x1 ehlnat_1p_x2 ehrnat_p1_y0
ehrnat_p1_y1 ehrnat_p1_z0 ehrnat_p1_z1
wlrnat_V_x_y wlrnat_V_x_z
forall (b c d e f : X) (wlx0 x0 wrx0 : a = b)
(wlx1 x1 wrx1 : c = d) (wlx2 x2 wrx2 : e = f)
(wly0 y0 wry0 : b = d) (wly1 y1 wry1 : a = c)
(wlz0 z0 wrz0 : d = f) (wlz1 z1 wrz1 : c = e)
(wlyz0 wryz0 : b = f) (wlyz1 wryz1 : a = e)
(ulnat_x0 : wlx0 @ 1 = 1 @ x0)
(urnat_x0 : wrx0 @ 1 = 1 @ x0)
(ulnat_x1 : wlx1 @ 1 = 1 @ x1)
(urnat_x1 : wrx1 @ 1 = 1 @ x1)
(ulnat_x2 : wlx2 @ 1 = 1 @ x2)
(urnat_x2 : wrx2 @ 1 = 1 @ x2)
(ulnat_y0 : wly0 @ 1 = 1 @ y0)
(urnat_y0 : wry0 @ 1 = 1 @ y0)
(ulnat_y1 : wly1 @ 1 = 1 @ y1)
(urnat_y1 : wry1 @ 1 = 1 @ y1)
(ulnat_z0 : wlz0 @ 1 = 1 @ z0)
(urnat_z0 : wrz0 @ 1 = 1 @ z0)
(ulnat_z1 : wlz1 @ 1 = 1 @ z1)
(urnat_z1 : wrz1 @ 1 = 1 @ z1)
(ulnat_yz0 : wlyz0 @ 1 = 1 @ (y0 @ z0))
(urnat_yz0 : wryz0 @ 1 = 1 @ (y0 @ z0))
(ulnat_yz1 : wlyz1 @ 1 = 1 @ (y1 @ z1))
(urnat_yz1 : wryz1 @ 1 = 1 @ (y1 @ z1))
(ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0)
(ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1)
(ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2)
(ehrnat_y0 : wry0 @ 1 = 1 @ wly0)
(ehrnat_y1 : wry1 @ 1 = 1 @ wly1)
(ehrnat_z0 : wrz0 @ 1 = 1 @ wlz0)
(ehrnat_z1 : wrz1 @ 1 = 1 @ wlz1)
(ehrnat_yz0 : wryz0 @ 1 = 1 @ wlyz0)
(ehrnat_yz1 : wryz1 @ 1 = 1 @ wlyz1)
(wlrnat_x_y : wlx0 @ wry0 = wry1 @ wlx1)
(wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0)
(wlrnat_x_z : wlx1 @ wrz0 = wrz1 @ wlx2)
(wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0)
(wlrnat_x_yz : wlx0 @ wryz0 = wryz1 @ wlx2)
(wlrnat_yz_x : wlyz1 @ wrx2 = wrx0 @ wlyz0)
(wrpp_yz0 : wry0 @ wrz0 = wryz0)
(wlpp_yz0 : wly0 @ wlz0 = wlyz0)
(wrpp_yz1 : wry1 @ wrz1 = wryz1)
(wlpp_yz1 : wly1 @ wlz1 = wlyz1)
(H_ulnat_yz0 : ulnat_y0 [-] ulnat_z0 =
whiskerR wlpp_yz0 1 @ ulnat_yz0)
(H_urnat_yz0 : urnat_y0 [-] urnat_z0 =
whiskerR wrpp_yz0 1 @ urnat_yz0)
(H_ulnat_yz1 : ulnat_y1 [-] ulnat_z1 =
whiskerR wlpp_yz1 1 @ ulnat_yz1)
(H_urnat_yz1 : urnat_y1 [-] urnat_z1 =
whiskerR wrpp_yz1 1 @ urnat_yz1)
(H_ehrnat_yz0 : (ehrnat_y0 [-] ehrnat_z0) @
whiskerL 1 wlpp_yz0 =
whiskerR wrpp_yz0 1 @ ehrnat_yz0)
(H_ehrnat_yz1 : (ehrnat_y1 [-] ehrnat_z1) @
whiskerL 1 wlpp_yz1 =
whiskerR wrpp_yz1 1 @ ehrnat_yz1)
(H_wlrnat_x_yz : (wlrnat_x_y [I] wlrnat_x_z) @
whiskerR wrpp_yz1 wlx2 =
whiskerL wlx0 wrpp_yz0 @ wlrnat_x_yz)
(H_wlrnat_yz_x : (wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 wlpp_yz0 =
whiskerR wlpp_yz1 wrx2 @ wlrnat_yz_x)
(ehlnat_1p_x0 : (ehlnat_x0 [I] urnat_x0) @ 1 =
1 @ ulnat_x0)
(ehlnat_1p_x1 : (ehlnat_x1 [I] urnat_x1) @ 1 =
1 @ ulnat_x1)
(ehlnat_1p_x2 : (ehlnat_x2 [I] urnat_x2) @ 1 =
1 @ ulnat_x2)
(ehrnat_p1_y0 : (ehrnat_y0 [I] ulnat_y0) @ 1 =
1 @ urnat_y0)
(ehrnat_p1_y1 : (ehrnat_y1 [I] ulnat_y1) @ 1 =
1 @ urnat_y1)
(ehrnat_p1_z0 : (ehrnat_z0 [I] ulnat_z0) @ 1 =
1 @ urnat_z0)
(ehrnat_p1_z1 : (ehrnat_z1 [I] ulnat_z1) @ 1 =
1 @ urnat_z1)
(wlrnat_V_x_y : whiskerR wlrnat_x_y 1 @
(ehrnat_y1 [-] ehlnat_x1) =
(ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^)
(wlrnat_V_x_z : whiskerR wlrnat_x_z 1 @
(ehrnat_z1 [-] ehlnat_x2) =
(ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^),
let eh_x_y :=
concat_p_pp x0 y0 z0 @
whiskerR
(((rlucancel_inv (ulnat_x0 [-] urnat_y0))^ @
wlrnat_x_y) @
rlucancel_inv (urnat_y1 [-] ulnat_x1)) z0 in
(((whiskerR
(concat_p1
(((rlucancel_inv (ulnat_x0 [-] urnat_yz0))^ @
wlrnat_x_yz) @
rlucancel_inv (urnat_yz1 [-] ulnat_x2)) @@
concat_p1
(((rlucancel_inv (ulnat_yz1 [-] urnat_x2))^ @
wlrnat_yz_x) @
rlucancel_inv (urnat_x0 [-] ulnat_yz0)))
eh_x_y @
whiskerR
(eh_V_gen ehlnat_1p_x0 ehlnat_1p_x2
(Ehrnat_p1_pp 1 1 H_ehrnat_yz0 H_ulnat_yz0
H_urnat_yz0 ehrnat_p1_y0 ehrnat_p1_z0)
(Ehrnat_p1_pp 1 1 H_ehrnat_yz1 H_ulnat_yz1
H_urnat_yz1 ehrnat_p1_y1 ehrnat_p1_z1)
(Wlrnat_V_p_pp H_ehrnat_yz0 H_ehrnat_yz1
H_wlrnat_x_yz H_wlrnat_yz_x wlrnat_V_x_y
wlrnat_V_x_z)) eh_x_y) @
lrucancel 1 ) @
whiskerL eh_x_y (concat_pp_p_p_pp y1 x1 z0)^) @
whiskerL eh_x_y
(concat_p1 (concat_pp_p y1 x1 z0) @@
concat_p1 (concat_p_pp y1 x1 z0))^ =
(eh_p_pp_gen H_urnat_yz0 H_urnat_yz1 H_wlrnat_x_yz
[-] lrucancel
(whiskerL (concat_pp_p y1 z1 x2)
(ap
(fun p : z1 @ x2 = x1 @ z0 =>
whiskerL y1 p)
(moveL_V1
(((rlucancel_inv
(ulnat_z1 [-] urnat_x2))^ @
wlrnat_z_x) @
rlucancel_inv (urnat_x1 [-] ulnat_z0))
(((rlucancel_inv
(ulnat_x1 [-] urnat_z0))^ @
wlrnat_x_z) @
rlucancel_inv (urnat_z1 [-] ulnat_x2))
(eh_V_gen ehlnat_1p_x1 ehlnat_1p_x2
ehrnat_p1_z0 ehrnat_p1_z1
wlrnat_V_x_z)))))
[-] (eh_pp_p_gen H_ulnat_yz1 H_ulnat_yz0 H_wlrnat_yz_x
[-] lrucancel
(whiskerL (concat_p_pp x0 y0 z0)
(ap
(fun p : x0 @ y0 = y1 @ x1 =>
whiskerR p z0)
(moveL_1V
(((rlucancel_inv
(ulnat_x0 [-] urnat_y0))^ @
wlrnat_x_y) @
rlucancel_inv
(urnat_y1 [-] ulnat_x1))
(((rlucancel_inv
(ulnat_y1 [-] urnat_x1))^ @
wlrnat_y_x) @
rlucancel_inv
(urnat_x0 [-] ulnat_y0))
(eh_V_gen ehlnat_1p_x0
ehlnat_1p_x1 ehrnat_p1_y0
ehrnat_p1_y1 wlrnat_V_x_y)))))
apply lem.X : Type a : X lem : forall
P : forall (x x0 x1 x2 x3 : X)
(x4 x5 x6 : a = x) (x7 x8 x9 : x0 = x1)
(x10 x11 x12 : x2 = x3)
(x13 x14 x15 : x = x1)
(x16 x17 x18 : a = x0)
(x19 x20 x21 : x1 = x3)
(x22 x23 x24 : x0 = x2)
(x25 x26 : x = x3) (x27 x28 : a = x2)
(x29 : x4 @ 1 = 1 @ x5)
(x30 : x6 @ 1 = 1 @ x5)
(x31 : x7 @ 1 = 1 @ x8)
(x32 : x9 @ 1 = 1 @ x8)
(x33 : x10 @ 1 = 1 @ x11)
(x34 : x12 @ 1 = 1 @ x11)
(x35 : x13 @ 1 = 1 @ x14)
(x36 : x15 @ 1 = 1 @ x14)
(x37 : x16 @ 1 = 1 @ x17)
(x38 : x18 @ 1 = 1 @ x17)
(x39 : x19 @ 1 = 1 @ x20)
(x40 : x21 @ 1 = 1 @ x20)
(x41 : x22 @ 1 = 1 @ x23)
(x42 : x24 @ 1 = 1 @ x23)
(x43 : x25 @ 1 = 1 @ (x14 @ x20))
(x44 : x26 @ 1 = 1 @ (x14 @ x20))
(x45 : x27 @ 1 = 1 @ (x17 @ x23))
(x46 : x28 @ 1 = 1 @ (x17 @ x23))
(x47 : x4 @ 1 = 1 @ x6)
(x48 : x7 @ 1 = 1 @ x9)
(x49 : x10 @ 1 = 1 @ x12)
(x50 : x15 @ 1 = 1 @ x13)
(x51 : x18 @ 1 = 1 @ x16)
(x52 : x21 @ 1 = 1 @ x19)
(x53 : x24 @ 1 = 1 @ x22)
(x54 : x26 @ 1 = 1 @ x25)
(x55 : x28 @ 1 = 1 @ x27)
(x56 : x4 @ x15 = x18 @ x7)
(x57 : x16 @ x9 = x6 @ x13)
(x58 : x7 @ x21 = x24 @ x10)
(x59 : x22 @ x12 = x9 @ x19)
(x60 : x4 @ x26 = x28 @ x10)
(x61 : x27 @ x12 = x6 @ x25)
(x62 : x15 @ x21 = x26)
(x63 : x13 @ x19 = x25)
(x64 : x18 @ x24 = x28)
(x65 : x16 @ x22 = x27),
x35 [-] x39 = whiskerR x63 1 @ x43 ->
x36 [-] x40 = whiskerR x62 1 @ x44 ->
x37 [-] x41 = whiskerR x65 1 @ x45 ->
x38 [-] x42 = whiskerR x64 1 @ x46 ->
(x50 [-] x52) @ whiskerL 1 x63 =
whiskerR x62 1 @ x54 ->
(x51 [-] x53) @ whiskerL 1 x65 =
whiskerR x64 1 @ x55 ->
(x56 [I] x58) @ whiskerR x64 x10 =
whiskerL x4 x62 @ x60 ->
(x57 [-] x59) @ whiskerL x6 x63 =
whiskerR x65 x12 @ x61 ->
(x47 [I] x30) @ 1 = 1 @ x29 ->
(x48 [I] x32) @ 1 = 1 @ x31 ->
(x49 [I] x34) @ 1 = 1 @ x33 ->
(x50 [I] x35) @ 1 = 1 @ x36 ->
(x51 [I] x37) @ 1 = 1 @ x38 ->
(x52 [I] x39) @ 1 = 1 @ x40 ->
(x53 [I] x41) @ 1 = 1 @ x42 ->
whiskerR x56 1 @ (x51 [-] x48) =
(x47 [-] x50) @ whiskerL 1 x57^ ->
whiskerR x58 1 @ (x53 [-] x49) =
(x48 [-] x52) @ whiskerL 1 x59^ -> Type ,
P a a a a a 1 1 1 1 1 1 1 1 1 1 1 1 1 1 1 1 1 1
1 1 1 (1 @ 1 ) (1 @ 1 ) (1 @ 1 )
(1 @ 1 ) (rlucancel 1 [I] rlucancel 1 )
(rlucancel 1 ) (rlucancel 1 [I] rlucancel 1 )
(rlucancel 1 ) (rlucancel 1 [I] rlucancel 1 )
(rlucancel 1 ) (rlucancel 1 )
(rlucancel 1 [I] rlucancel 1 )
(rlucancel 1 ) (rlucancel 1 [I] rlucancel 1 )
(rlucancel 1 ) (rlucancel 1 [I] rlucancel 1 )
(rlucancel 1 ) (rlucancel 1 [I] rlucancel 1 )
((whiskerR 1 1 )^ @
(rlucancel 1 [-] rlucancel 1 ))
((whiskerR 1 1 )^ @
((rlucancel 1 [I] rlucancel 1 )
[-] (rlucancel 1 [I] rlucancel 1 )))
((whiskerR 1 1 )^ @
(rlucancel 1 [-] rlucancel 1 ))
((whiskerR 1 1 )^ @
((rlucancel 1 [I] rlucancel 1 )
[-] (rlucancel 1 [I] rlucancel 1 )))
(rlucancel 1 ) (rlucancel 1 )
(rlucancel 1 ) (rlucancel 1 )
(rlucancel 1 ) (rlucancel 1 )
(rlucancel 1 )
((whiskerR 1 1 )^ @
((rlucancel 1 [-] rlucancel 1 ) @ whiskerL 1 1 ))
((whiskerR 1 1 )^ @
((rlucancel 1 [-] rlucancel 1 ) @ whiskerL 1 1 ))
(((concat_p1 (1 @ 1 ))^ @
(((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 (lrucancel 1 )^) @
(rlucancel 1 [-] rlucancel 1 )^)) @
concat_p1 (1 @ 1 )) (lrucancel 1 )
(((concat_p1 (1 @ 1 ))^ @
(((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 (lrucancel 1 )^) @
(rlucancel 1 [-] rlucancel 1 )^)) @
concat_p1 (1 @ 1 )) (lrucancel 1 )
((whiskerL 1 1 )^ @
((((concat_p1 (1 @ 1 ))^ @
(((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 (lrucancel 1 )^) @
(rlucancel 1 [-] rlucancel 1 )^)) @
concat_p1 (1 @ 1 )
[I] ((concat_p1 (1 @ 1 ))^ @
(((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 (lrucancel 1 )^) @
(rlucancel 1 [-] rlucancel 1 )^)) @
concat_p1 (1 @ 1 )) @
whiskerR 1 1 ))
((whiskerR 1 1 )^ @
((lrucancel 1 [-] lrucancel 1 ) @ whiskerL 1 1 ))
1 1 1 1
(equiv_moveL_Mp
((whiskerR 1 1 )^ @
(rlucancel 1 [-] rlucancel 1 ))
(rlucancel 1 [-] rlucancel 1 )
(whiskerR 1 1 ) 1 )
(equiv_moveL_Mp
((whiskerR 1 1 )^ @
((rlucancel 1 [I] rlucancel 1 )
[-] (rlucancel 1 [I] rlucancel 1 )))
((rlucancel 1 [I] rlucancel 1 )
[-] (rlucancel 1 [I] rlucancel 1 ))
(whiskerR 1 1 ) 1 )
(equiv_moveL_Mp
((whiskerR 1 1 )^ @
(rlucancel 1 [-] rlucancel 1 ))
(rlucancel 1 [-] rlucancel 1 )
(whiskerR 1 1 ) 1 )
(equiv_moveL_Mp
((whiskerR 1 1 )^ @
((rlucancel 1 [I] rlucancel 1 )
[-] (rlucancel 1 [I] rlucancel 1 )))
((rlucancel 1 [I] rlucancel 1 )
[-] (rlucancel 1 [I] rlucancel 1 ))
(whiskerR 1 1 ) 1 )
(equiv_moveL_Mp
((whiskerR 1 1 )^ @
((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 1 ))
((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 1 ) (whiskerR 1 1 ) 1 )
(equiv_moveL_Mp
((whiskerR 1 1 )^ @
((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 1 ))
((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 1 ) (whiskerR 1 1 ) 1 )
(equiv_moveL_Mp
((whiskerL 1 1 )^ @
((((concat_p1 (1 @ 1 ))^ @
(((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 (lrucancel 1 )^) @
(rlucancel 1 [-] rlucancel 1 )^)) @
concat_p1 (1 @ 1 )
[I] ((concat_p1 (1 @ 1 ))^ @
(((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 (lrucancel 1 )^) @
(rlucancel 1 [-] rlucancel 1 )^)) @
concat_p1 (1 @ 1 )) @
whiskerR 1 1 ))
((((concat_p1 (1 @ 1 ))^ @
(((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 (lrucancel 1 )^) @
(rlucancel 1 [-] rlucancel 1 )^)) @
concat_p1 (1 @ 1 )
[I] ((concat_p1 (1 @ 1 ))^ @
(((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 (lrucancel 1 )^) @
(rlucancel 1 [-] rlucancel 1 )^)) @
concat_p1 (1 @ 1 )) @
whiskerR 1 1 ) (whiskerL 1 1 ) 1 )
(equiv_moveL_Mp
((whiskerR 1 1 )^ @
((lrucancel 1 [-] lrucancel 1 ) @
whiskerL 1 1 ))
((lrucancel 1 [-] lrucancel 1 ) @
whiskerL 1 1 ) (whiskerR 1 1 ) 1 )
(rlucancel 1 ) (rlucancel 1 )
(rlucancel 1 ) (rlucancel 1 )
(rlucancel 1 ) (rlucancel 1 )
(rlucancel 1 )
(equiv_helper (rlucancel 1 [-] rlucancel 1 )
((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 (lrucancel 1 )^)
(((concat_p1 (1 @ 1 ))^ @
(((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 (lrucancel 1 )^) @
(rlucancel 1 [-] rlucancel 1 )^)) @
concat_p1 (1 @ 1 )) 1 )
(equiv_helper (rlucancel 1 [-] rlucancel 1 )
((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 (lrucancel 1 )^)
(((concat_p1 (1 @ 1 ))^ @
(((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 (lrucancel 1 )^) @
(rlucancel 1 [-] rlucancel 1 )^)) @
concat_p1 (1 @ 1 )) 1 ) ->
forall (b c d e f : X) (wlx0 x0 wrx0 : a = b)
(wlx1 x1 wrx1 : c = d) (wlx2 x2 wrx2 : e = f)
(wly0 y0 wry0 : b = d) (wly1 y1 wry1 : a = c)
(wlz0 z0 wrz0 : d = f) (wlz1 z1 wrz1 : c = e)
(wlyz0 wryz0 : b = f) (wlyz1 wryz1 : a = e)
(ulnat_x0 : wlx0 @ 1 = 1 @ x0)
(urnat_x0 : wrx0 @ 1 = 1 @ x0)
(ulnat_x1 : wlx1 @ 1 = 1 @ x1)
(urnat_x1 : wrx1 @ 1 = 1 @ x1)
(ulnat_x2 : wlx2 @ 1 = 1 @ x2)
(urnat_x2 : wrx2 @ 1 = 1 @ x2)
(ulnat_y0 : wly0 @ 1 = 1 @ y0)
(urnat_y0 : wry0 @ 1 = 1 @ y0)
(ulnat_y1 : wly1 @ 1 = 1 @ y1)
(urnat_y1 : wry1 @ 1 = 1 @ y1)
(ulnat_z0 : wlz0 @ 1 = 1 @ z0)
(urnat_z0 : wrz0 @ 1 = 1 @ z0)
(ulnat_z1 : wlz1 @ 1 = 1 @ z1)
(urnat_z1 : wrz1 @ 1 = 1 @ z1)
(ulnat_yz0 : wlyz0 @ 1 = 1 @ (y0 @ z0))
(urnat_yz0 : wryz0 @ 1 = 1 @ (y0 @ z0))
(ulnat_yz1 : wlyz1 @ 1 = 1 @ (y1 @ z1))
(urnat_yz1 : wryz1 @ 1 = 1 @ (y1 @ z1))
(ehlnat_x0 : wlx0 @ 1 = 1 @ wrx0)
(ehlnat_x1 : wlx1 @ 1 = 1 @ wrx1)
(ehlnat_x2 : wlx2 @ 1 = 1 @ wrx2)
(ehrnat_y0 : wry0 @ 1 = 1 @ wly0)
(ehrnat_y1 : wry1 @ 1 = 1 @ wly1)
(ehrnat_z0 : wrz0 @ 1 = 1 @ wlz0)
(ehrnat_z1 : wrz1 @ 1 = 1 @ wlz1)
(ehrnat_yz0 : wryz0 @ 1 = 1 @ wlyz0)
(ehrnat_yz1 : wryz1 @ 1 = 1 @ wlyz1)
(wlrnat_x_y : wlx0 @ wry0 = wry1 @ wlx1)
(wlrnat_y_x : wly1 @ wrx1 = wrx0 @ wly0)
(wlrnat_x_z : wlx1 @ wrz0 = wrz1 @ wlx2)
(wlrnat_z_x : wlz1 @ wrx2 = wrx1 @ wlz0)
(wlrnat_x_yz : wlx0 @ wryz0 = wryz1 @ wlx2)
(wlrnat_yz_x : wlyz1 @ wrx2 = wrx0 @ wlyz0)
(wrpp_yz0 : wry0 @ wrz0 = wryz0)
(wlpp_yz0 : wly0 @ wlz0 = wlyz0)
(wrpp_yz1 : wry1 @ wrz1 = wryz1)
(wlpp_yz1 : wly1 @ wlz1 = wlyz1)
(H_ulnat_yz0 : ulnat_y0 [-] ulnat_z0 =
whiskerR wlpp_yz0 1 @ ulnat_yz0)
(H_urnat_yz0 : urnat_y0 [-] urnat_z0 =
whiskerR wrpp_yz0 1 @ urnat_yz0)
(H_ulnat_yz1 : ulnat_y1 [-] ulnat_z1 =
whiskerR wlpp_yz1 1 @ ulnat_yz1)
(H_urnat_yz1 : urnat_y1 [-] urnat_z1 =
whiskerR wrpp_yz1 1 @ urnat_yz1)
(H_ehrnat_yz0 : (ehrnat_y0 [-] ehrnat_z0) @
whiskerL 1 wlpp_yz0 =
whiskerR wrpp_yz0 1 @ ehrnat_yz0)
(H_ehrnat_yz1 : (ehrnat_y1 [-] ehrnat_z1) @
whiskerL 1 wlpp_yz1 =
whiskerR wrpp_yz1 1 @ ehrnat_yz1)
(H_wlrnat_x_yz : (wlrnat_x_y [I] wlrnat_x_z) @
whiskerR wrpp_yz1 wlx2 =
whiskerL wlx0 wrpp_yz0 @
wlrnat_x_yz)
(H_wlrnat_yz_x : (wlrnat_y_x [-] wlrnat_z_x) @
whiskerL wrx0 wlpp_yz0 =
whiskerR wlpp_yz1 wrx2 @
wlrnat_yz_x)
(ehlnat_1p_x0 : (ehlnat_x0 [I] urnat_x0) @ 1 =
1 @ ulnat_x0)
(ehlnat_1p_x1 : (ehlnat_x1 [I] urnat_x1) @ 1 =
1 @ ulnat_x1)
(ehlnat_1p_x2 : (ehlnat_x2 [I] urnat_x2) @ 1 =
1 @ ulnat_x2)
(ehrnat_p1_y0 : (ehrnat_y0 [I] ulnat_y0) @ 1 =
1 @ urnat_y0)
(ehrnat_p1_y1 : (ehrnat_y1 [I] ulnat_y1) @ 1 =
1 @ urnat_y1)
(ehrnat_p1_z0 : (ehrnat_z0 [I] ulnat_z0) @ 1 =
1 @ urnat_z0)
(ehrnat_p1_z1 : (ehrnat_z1 [I] ulnat_z1) @ 1 =
1 @ urnat_z1)
(wlrnat_V_x_y : whiskerR wlrnat_x_y 1 @
(ehrnat_y1 [-] ehlnat_x1) =
(ehlnat_x0 [-] ehrnat_y0) @
whiskerL 1 wlrnat_y_x^)
(wlrnat_V_x_z : whiskerR wlrnat_x_z 1 @
(ehrnat_z1 [-] ehlnat_x2) =
(ehlnat_x1 [-] ehrnat_z0) @
whiskerL 1 wlrnat_z_x^),
P b c d e f wlx0 x0 wrx0 wlx1 x1 wrx1 wlx2 x2
wrx2 wly0 y0 wry0 wly1 y1 wry1 wlz0 z0 wrz0
wlz1 z1 wrz1 wlyz0 wryz0 wlyz1 wryz1 ulnat_x0
urnat_x0 ulnat_x1 urnat_x1 ulnat_x2 urnat_x2
ulnat_y0 urnat_y0 ulnat_y1 urnat_y1 ulnat_z0
urnat_z0 ulnat_z1 urnat_z1 ulnat_yz0 urnat_yz0
ulnat_yz1 urnat_yz1 ehlnat_x0 ehlnat_x1
ehlnat_x2 ehrnat_y0 ehrnat_y1 ehrnat_z0
ehrnat_z1 ehrnat_yz0 ehrnat_yz1 wlrnat_x_y
wlrnat_y_x wlrnat_x_z wlrnat_z_x wlrnat_x_yz
wlrnat_yz_x wrpp_yz0 wlpp_yz0 wrpp_yz1
wlpp_yz1 H_ulnat_yz0 H_urnat_yz0 H_ulnat_yz1
H_urnat_yz1 H_ehrnat_yz0 H_ehrnat_yz1
H_wlrnat_x_yz H_wlrnat_yz_x ehlnat_1p_x0
ehlnat_1p_x1 ehlnat_1p_x2 ehrnat_p1_y0
ehrnat_p1_y1 ehrnat_p1_z0 ehrnat_p1_z1
wlrnat_V_x_y wlrnat_V_x_z
(((whiskerR
(concat_p1
(((rlucancel_inv
((rlucancel 1 [I] rlucancel 1 )
[-] (whiskerR 1 1 )^ @
((rlucancel 1 [I] rlucancel 1 )
[-] (rlucancel 1 [I] rlucancel 1 ))))^ @
((whiskerL 1 1 )^ @
((((concat_p1 ...)^ @ (... @ ...^)) @
concat_p1 (1 @ 1 )
[I] ((concat_p1 ...)^ @ (... @ ...^)) @
concat_p1 (1 @ 1 )) @
whiskerR 1 1 ))) @
rlucancel_inv
((whiskerR 1 1 )^ @
((rlucancel 1 [I] rlucancel 1 )
[-] (rlucancel 1 [I] rlucancel 1 ))
[-] (rlucancel 1 [I] rlucancel 1 ))) @@
concat_p1
(((rlucancel_inv
((whiskerR 1 1 )^ @
(rlucancel 1 [-] rlucancel 1 )
[-] rlucancel 1 ))^ @
((whiskerR 1 1 )^ @
((lrucancel 1 [-] lrucancel 1 ) @
whiskerL 1 1 ))) @
rlucancel_inv
(rlucancel 1
[-] (whiskerR 1 1 )^ @
(rlucancel 1 [-] rlucancel 1 ))))
(concat_p_pp 1 1 1 @
whiskerR
(((rlucancel_inv
((rlucancel 1 [I] rlucancel 1 )
[-] (rlucancel 1 [I] rlucancel 1 )))^ @
(((concat_p1 (1 @ 1 ))^ @
(((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 (lrucancel 1 )^) @
(rlucancel 1 [-] rlucancel 1 )^)) @
concat_p1 (1 @ 1 ))) @
rlucancel_inv
((rlucancel 1 [I] rlucancel 1 )
[-] (rlucancel 1 [I] rlucancel 1 ))) 1 ) @
whiskerR
(eh_V_gen (rlucancel 1 )
(rlucancel 1 )
(Ehrnat_p1_pp 1 1
(equiv_moveL_Mp
((whiskerR 1 1 )^ @
((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 1 ))
((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 1 )
(whiskerR 1 1 ) 1 )
(equiv_moveL_Mp
((whiskerR 1 1 )^ @
(rlucancel 1 [-] rlucancel 1 ))
(rlucancel 1 [-] rlucancel 1 )
(whiskerR 1 1 ) 1 )
(equiv_moveL_Mp
((whiskerR 1 1 )^ @
((rlucancel 1 [I] rlucancel 1 )
[-] (rlucancel 1 [I] rlucancel 1 )))
((rlucancel 1 [I] rlucancel 1 )
[-] (rlucancel 1 [I] rlucancel 1 ))
(whiskerR 1 1 ) 1 )
(rlucancel 1 )
(rlucancel 1 ))
(Ehrnat_p1_pp 1 1
(equiv_moveL_Mp
((whiskerR 1 1 )^ @
((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 1 ))
((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 1 )
(whiskerR 1 1 ) 1 )
(equiv_moveL_Mp
((whiskerR 1 1 )^ @
(rlucancel 1 [-] rlucancel 1 ))
(rlucancel 1 [-] rlucancel 1 )
(whiskerR 1 1 ) 1 )
(equiv_moveL_Mp
((whiskerR 1 1 )^ @
((rlucancel 1 [I] rlucancel 1 )
[-] (rlucancel 1 [I] rlucancel 1 )))
((rlucancel 1 [I] rlucancel 1 )
[-] (rlucancel 1 [I] rlucancel 1 ))
(whiskerR 1 1 ) 1 )
(rlucancel 1 )
(rlucancel 1 ))
(Wlrnat_V_p_pp
(equiv_moveL_Mp
((whiskerR 1 1 )^ @
((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 1 ))
((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 1 )
(whiskerR 1 1 ) 1 )
(equiv_moveL_Mp
((whiskerR 1 1 )^ @
((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 1 ))
((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 1 )
(whiskerR 1 1 ) 1 )
(equiv_moveL_Mp
((whiskerL 1 1 )^ @
((((concat_p1 (...))^ @
((...) @ (...)^)) @
concat_p1 (1 @ 1 )
[I] ((concat_p1 (...))^ @
((...) @ (...)^)) @
concat_p1 (1 @ 1 )) @
whiskerR 1 1 ))
((((concat_p1 (1 @ 1 ))^ @
(((...) @ whiskerL 1 ...^) @
(rlucancel 1 [-] rlucancel 1 )^)) @
concat_p1 (1 @ 1 )
[I] ((concat_p1 (1 @ 1 ))^ @
(((...) @ whiskerL 1 ...^) @
(rlucancel 1 [-] rlucancel 1 )^)) @
concat_p1 (1 @ 1 )) @
whiskerR 1 1 )
(whiskerL 1 1 ) 1 )
(equiv_moveL_Mp
((whiskerR 1 1 )^ @
((lrucancel 1 [-] lrucancel 1 ) @
whiskerL 1 1 ))
((lrucancel 1 [-] lrucancel 1 ) @
whiskerL 1 1 )
(whiskerR 1 1 ) 1 )
(equiv_helper
(rlucancel 1 [-] rlucancel 1 )
((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 (lrucancel 1 )^)
(((concat_p1 (1 @ 1 ))^ @
(((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 (lrucancel 1 )^) @
(rlucancel 1 [-] rlucancel 1 )^)) @
concat_p1 (1 @ 1 )) 1 )
(equiv_helper
(rlucancel 1 [-] rlucancel 1 )
((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 (lrucancel 1 )^)
(((concat_p1 (1 @ 1 ))^ @
(((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 (lrucancel 1 )^) @
(rlucancel 1 [-] rlucancel 1 )^)) @
concat_p1 (1 @ 1 )) 1 )))
(concat_p_pp 1 1 1 @
whiskerR
(((rlucancel_inv
((rlucancel 1 [I] rlucancel 1 )
[-] (rlucancel 1 [I] rlucancel 1 )))^ @
(((concat_p1 (1 @ 1 ))^ @
(((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 (lrucancel 1 )^) @
(rlucancel 1 [-] rlucancel 1 )^)) @
concat_p1 (1 @ 1 ))) @
rlucancel_inv
((rlucancel 1 [I] rlucancel 1 )
[-] (rlucancel 1 [I] rlucancel 1 ))) 1 )) @
lrucancel 1 ) @
whiskerL
(concat_p_pp 1 1 1 @
whiskerR
(((rlucancel_inv
((rlucancel 1 [I] rlucancel 1 )
[-] (rlucancel 1 [I] rlucancel 1 )))^ @
(((concat_p1 (1 @ 1 ))^ @
(((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 (lrucancel 1 )^) @
(rlucancel 1 [-] rlucancel 1 )^)) @
concat_p1 (1 @ 1 ))) @
rlucancel_inv
((rlucancel 1 [I] rlucancel 1 )
[-] (rlucancel 1 [I] rlucancel 1 ))) 1 )
(concat_pp_p_p_pp 1 1 1 )^) @
whiskerL
(concat_p_pp 1 1 1 @
whiskerR
(((rlucancel_inv
((rlucancel 1 [I] rlucancel 1 )
[-] (rlucancel 1 [I] rlucancel 1 )))^ @
(((concat_p1 (1 @ 1 ))^ @
(((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 (lrucancel 1 )^) @
(rlucancel 1 [-] rlucancel 1 )^)) @
concat_p1 (1 @ 1 ))) @
rlucancel_inv
((rlucancel 1 [I] rlucancel 1 )
[-] (rlucancel 1 [I] rlucancel 1 ))) 1 )
(concat_p1 (concat_pp_p 1 1 1 ) @@
concat_p1 (concat_p_pp 1 1 1 ))^ =
(eh_p_pp_gen
(equiv_moveL_Mp
((whiskerR 1 1 )^ @
((rlucancel 1 [I] rlucancel 1 )
[-] (rlucancel 1 [I] rlucancel 1 )))
((rlucancel 1 [I] rlucancel 1 )
[-] (rlucancel 1 [I] rlucancel 1 ))
(whiskerR 1 1 ) 1 )
(equiv_moveL_Mp
((whiskerR 1 1 )^ @
((rlucancel 1 [I] rlucancel 1 )
[-] (rlucancel 1 [I] rlucancel 1 )))
((rlucancel 1 [I] rlucancel 1 )
[-] (rlucancel 1 [I] rlucancel 1 ))
(whiskerR 1 1 ) 1 )
(equiv_moveL_Mp
((whiskerL 1 1 )^ @
((((concat_p1 (1 @ 1 ))^ @
(((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 (lrucancel 1 )^) @
(rlucancel 1 [-] rlucancel 1 )^)) @
concat_p1 (1 @ 1 )
[I] ((concat_p1 (1 @ 1 ))^ @
(((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 (lrucancel 1 )^) @
(rlucancel 1 [-] rlucancel 1 )^)) @
concat_p1 (1 @ 1 )) @
whiskerR 1 1 ))
((((concat_p1 (1 @ 1 ))^ @
(((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 (lrucancel 1 )^) @
(rlucancel 1 [-] rlucancel 1 )^)) @
concat_p1 (1 @ 1 )
[I] ((concat_p1 (1 @ 1 ))^ @
(((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 (lrucancel 1 )^) @
(rlucancel 1 [-] rlucancel 1 )^)) @
concat_p1 (1 @ 1 )) @
whiskerR 1 1 ) (whiskerL 1 1 ) 1 )
[-] lrucancel
(whiskerL (concat_pp_p 1 1 1 )
(ap (fun p : 1 @ 1 = 1 @ 1 => whiskerL 1 p)
(moveL_V1
(((rlucancel_inv
(rlucancel 1 [-] rlucancel 1 ))^ @
lrucancel 1 ) @
rlucancel_inv
(rlucancel 1 [-] rlucancel 1 ))
(((rlucancel_inv
((rlucancel 1 [I] rlucancel 1 )
[-]
(rlucancel 1 [I] rlucancel 1 )))^ @
(((concat_p1 (1 @ 1 ))^ @
(((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 (lrucancel 1 )^) @
(rlucancel 1 [-] rlucancel 1 )^)) @
concat_p1 (1 @ 1 ))) @
rlucancel_inv
((rlucancel 1 [I] rlucancel 1 )
[-] (rlucancel 1 [I] rlucancel 1 )))
(eh_V_gen
(rlucancel 1 )
(rlucancel 1 )
(rlucancel 1 )
(rlucancel 1 )
(equiv_helper
(rlucancel 1 [-] rlucancel 1 )
((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 (lrucancel 1 )^)
(((concat_p1 (1 @ 1 ))^ @
(((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 (lrucancel 1 )^) @
(rlucancel 1 [-] rlucancel 1 )^)) @
concat_p1 (1 @ 1 )) 1 ))))))
[-] (eh_pp_p_gen
(equiv_moveL_Mp
((whiskerR 1 1 )^ @
(rlucancel 1 [-] rlucancel 1 ))
(rlucancel 1 [-] rlucancel 1 )
(whiskerR 1 1 ) 1 )
(equiv_moveL_Mp
((whiskerR 1 1 )^ @
(rlucancel 1 [-] rlucancel 1 ))
(rlucancel 1 [-] rlucancel 1 )
(whiskerR 1 1 ) 1 )
(equiv_moveL_Mp
((whiskerR 1 1 )^ @
((lrucancel 1 [-] lrucancel 1 ) @
whiskerL 1 1 ))
((lrucancel 1 [-] lrucancel 1 ) @
whiskerL 1 1 )
(whiskerR 1 1 ) 1 )
[-] lrucancel
(whiskerL (concat_p_pp 1 1 1 )
(ap
(fun p : 1 @ 1 = 1 @ 1 =>
whiskerR p 1 )
(moveL_1V
(((rlucancel_inv
((rlucancel 1 [I] rlucancel 1 )
[-]
(rlucancel 1 [I] rlucancel 1 )))^ @
(((concat_p1 (1 @ 1 ))^ @
(((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 (lrucancel 1 )^) @
(rlucancel 1 [-] rlucancel 1 )^)) @
concat_p1 (1 @ 1 ))) @
rlucancel_inv
((rlucancel 1 [I] rlucancel 1 )
[-]
(rlucancel 1 [I] rlucancel 1 )))
(((rlucancel_inv
(rlucancel 1 [-] rlucancel 1 ))^ @
lrucancel 1 ) @
rlucancel_inv
(rlucancel 1 [-] rlucancel 1 ))
(eh_V_gen
(rlucancel 1 )
(rlucancel 1 )
(rlucancel 1 )
(rlucancel 1 )
(equiv_helper
(rlucancel 1 [-] rlucancel 1 )
((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 (lrucancel 1 )^)
(((concat_p1 (1 @ 1 ))^ @
(((rlucancel 1 [-] rlucancel 1 ) @
whiskerL 1 (lrucancel 1 )^) @
(rlucancel 1 [-] rlucancel 1 )^)) @
concat_p1 (1 @ 1 )) 1 ))))))
reflexivity .
Qed .
Definition eh_V_p_pp {X } {a : X} (p q r : idpath (idpath a) = idpath (idpath a)) :
whiskerR (concat_p1 _ @@ concat_p1 _) _ @ whiskerR (eh_V p (q @ r)) _ @ lrucancel 1 @
whiskerL _ (Syllepsis.concat_pp_p_p_pp _ _ _)^ @ whiskerL _ (concat_p1 _ @@ concat_p1 _)^ =
(eh_p_pp_gen (urnat_pp q r) (urnat_pp q r) (wlrnat_p_pp p q r) [-]
lrucancel (whiskerL _ (ap (fun p => whiskerL q p) (moveL_V1 _ _ (eh_V p r))))) [-]
(eh_pp_p_gen (ulnat_pp q r) (ulnat_pp q r) (wlrnat_pp_p q r p) [-]
lrucancel (whiskerL _ (ap (fun p => whiskerR p r) (moveL_1V _ _ (eh_V p q))))).X : Type a : X p, q, r : 1 = 1
(((whiskerR
(concat_p1 (eh p (q @ r)) @@
concat_p1 (eh (q @ r) p))
(concat_p_pp p q r @ whiskerR (eh p q) r) @
whiskerR (eh_V p (q @ r))
(concat_p_pp p q r @ whiskerR (eh p q) r)) @
lrucancel 1 ) @
whiskerL (concat_p_pp p q r @ whiskerR (eh p q) r)
(concat_pp_p_p_pp q p r)^) @
whiskerL (concat_p_pp p q r @ whiskerR (eh p q) r)
(concat_p1 (concat_pp_p q p r) @@
concat_p1 (concat_p_pp q p r))^ =
(eh_p_pp_gen (urnat_pp q r) (urnat_pp q r)
(wlrnat_p_pp p q r)
[-] lrucancel
(whiskerL (concat_pp_p q r p)
(ap
(fun p0 : r @ p = p @ r => whiskerL q p0)
(moveL_V1 (eh r p) (eh p r) (eh_V p r)))))
[-] (eh_pp_p_gen (ulnat_pp q r) (ulnat_pp q r)
(wlrnat_pp_p q r p)
[-] lrucancel
(whiskerL (concat_p_pp p q r)
(ap
(fun p0 : p @ q = q @ p =>
whiskerR p0 r)
(moveL_1V (eh p q) (eh q p)
(eh_V p q)))))
Proof .X : Type a : X p, q, r : 1 = 1
(((whiskerR
(concat_p1 (eh p (q @ r)) @@
concat_p1 (eh (q @ r) p))
(concat_p_pp p q r @ whiskerR (eh p q) r) @
whiskerR (eh_V p (q @ r))
(concat_p_pp p q r @ whiskerR (eh p q) r)) @
lrucancel 1 ) @
whiskerL (concat_p_pp p q r @ whiskerR (eh p q) r)
(concat_pp_p_p_pp q p r)^) @
whiskerL (concat_p_pp p q r @ whiskerR (eh p q) r)
(concat_p1 (concat_pp_p q p r) @@
concat_p1 (concat_p_pp q p r))^ =
(eh_p_pp_gen (urnat_pp q r) (urnat_pp q r)
(wlrnat_p_pp p q r)
[-] lrucancel
(whiskerL (concat_pp_p q r p)
(ap
(fun p0 : r @ p = p @ r => whiskerL q p0)
(moveL_V1 (eh r p) (eh p r) (eh_V p r)))))
[-] (eh_pp_p_gen (ulnat_pp q r) (ulnat_pp q r)
(wlrnat_pp_p q r p)
[-] lrucancel
(whiskerL (concat_p_pp p q r)
(ap
(fun p0 : p @ q = q @ p =>
whiskerR p0 r)
(moveL_1V (eh p q) (eh q p)
(eh_V p q)))))
exact (eh_V_p_pp_gen _ _ _ (ehrnat_p1_pp q r) (ehrnat_p1_pp q r) (wlrnat_V_p_pp p q r)).
Defined .