This is formalization of Thierry Coquand's
cubical type theory.

Some of the main results proved so far are:
fill_from_comp_wf which proves that Kan filling operator
for cubical type can be constructed from "composition" operator.

pi_comp_wf2
which shows that from composition operators for and we can construct
composition operator for Π

sigma_comp_wf2
which shows that from composition operators for and we can construct
composition operator for Σ B

path_comp_wf
which shows that from composition operator for we can construct 
composition operator for (Path_A b)

glue-comp_wf2
which shows that from composition operators for and we can construct 
composition operator for Glue [psi ⊢→ (T;equiv-fun(f))] A.

case-type-comp_wf
shows that the "system type" constructed by compatible case has composition
operator. 

compU_wf
shows that the (infinite hierachy of) universes 
hava composition structure.

Using Glue and composition for Glue we can convert and
equivalence between types into path in the universe.
equiv-path_wf

(Univalence then follows from this.)


Comment: cubical_type_theory_basics
In this theory (by Thierry Coquand) the basic objects are:
1) Finite sets of names: names(I)  (we use natural numbers for the names).
    (see names_wf)
2) Morphisms: I ⟶ ==  names(J) ⟶ Point(dM(I)) 
    where ⌜dM(I)⌝ is the free De Morgan algebra generated by 
    ⌜<i>⌝ and ⌜<1-i>⌝ for ⌜i ∈ names(I)⌝
   (see dM_wf dM_inc_wf dM_opp_wf names-hom_wf)
3) With the appropriate composition,
       ∀[f:I ⟶ J]. ∀[g:J ⟶ K].  (g ⋅ f ∈ I ⟶ K),
    on the morphisms, these constitute
   the category of names CubeCat
   (see nh-comp_wf cube-cat_wf)

4) cubical set is presheaf over this category:
    CubicalSet ==  Functor(op-cat(names-cat);TypeCat)
   (see cubical_set_wf)

5) cubical set map A ⟶ B  between cubical sets A, is therefore a
   presheaf map, natural transformation 
          nat-trans(op-cat(names-cat);TypeCat;A;B)
  (see cube_set_map_wf)

The basic cubical sets are the following:

a) The formal cube 
    formal-cube(I) ==  <λJ.J ⟶ I, λI,J,f,g. g ⋅ f> 
    that is the representable presheaf defined by I.
   (see formal-cube_wf formal-cube-is-rep-pre-sheaf)

   If Γ is cubical set and ⌜ρ ∈ Γ(I)⌝ then ρ induces
   cubical set map <ρ> ∈ formal-cube(I) ⟶ Γ
   (see context-map_wf

b) The interval 𝕀 ==  <λI.Point(dM(I)), λJ,I,f. dM-lift(I;J;f)>
    whose objects are the finitely generated free De Morgan algebras.
   (see interval-presheaf_wf)

c) The "face presheaf"
       𝔽 ==  <λI.Point(face_lattice(I)), λJ,I,f. <f>>
   whose objects are the finitely generated "face lattices"
   which are free distributive lattices generated by 
   ⌜(i=0)⌝,⌜(i=1)⌝ (for ⌜i ∈ names(I)⌝subject to the relation
   ⌜(i=0) ∧ (i=1) 0⌝
   (see face-presheaf_wf)

d) member psi ∈ 𝔽(I) can be seen as formula about the points
   of the formal cube. This defines another cubical set
   I,psi (defined as rep-sub-sheaf(names-cat;I;λJ,f. (psi f) 1))
   which is the "cubical subset" of the formal cube satisfying the 
   formula psi.

   There is cubical set map iota ∈ I,psi ⟶ formal-cube(I) 
   that shows that I,psi  is subobject of formal-cube(I)

   Also, map ⌜f ∈ J ⟶ I⌝ induces cubical set map 
     subset-trans(I;J;f;psi) ∈ J,(psi)<f> ⟶ I,psi
   (see cubical-subset_wf subset-iota_wf subset-trans_wf)

e) trivial cubical set is the discrete cubical set 
    discrete-cube(A) ==  <λJ.A, λJ,K,f,x. x>
   that has type at every dimension and only the identity morphisms.
   special case is () ==  <λJ.Unit, λJ,K,f,x. ⋅>
   which represents the empty context in Martin Lof Type Theory.
   (see discrete-cube_wf trivial-cube-set_wf)


Comment: category of names doc
=========================== cube category =================================

The base category for the cubical set pre-sheaf is 
CubeCat ==  <fset(ℕ), λI,J. I ⟶ J, λI.1, λI,J,K,f,g. g ⋅ f> 

The objects are finite sets of names (we can use the natural numbers)
and the morphisms from to are maps 
I ⟶ ==  {n:ℕn ∈ J}  ⟶ Point(dM(I)).

dM(I) is the free de-Morgan algebra generated by the finite set I.
It is structure with type of "points" and meet, join, negation, and 1. 



Definition: names
names(I) ==  {n:ℕn ∈ I} 

Lemma: names_wf
[I:fset(ℕ)]. (names(I) ∈ Type)

Lemma: names-subtype
[I,J:fset(ℕ)].  names(I) ⊆names(J) supposing I ⊆ J

Definition: names-deq
NamesDeq ==  IntDeq

Lemma: names-deq_wf
[I:fset(ℕ)]. (NamesDeq ∈ EqDecider(names(I)))

Lemma: fset-names-to-list
[I:fset(ℕ)]. ∀s:fset(names(I)). ∃L:names(I) List. (L s ∈ fset(names(I)))

Definition: names-list
names-list(s) ==  fst((TERMOF{fset-names-to-list:o, 1:l} s))

Lemma: names-list_wf
[I:fset(ℕ)]. ∀[s:fset(names(I))].  (names-list(s) ∈ {L:names(I) List| s ∈ fset(names(I))} )

Definition: dM
dM(I) ==  free-DeMorgan-algebra(names(I);NamesDeq)

Lemma: dM_wf
[I:fset(ℕ)]. (dM(I) ∈ DeMorganAlgebra)

Lemma: dM-point
[I:Top]
  (Point(dM(I)) {ac:fset(fset(names(I) names(I)))| 
                   ↑fset-antichain(union-deq(names(I);names(I);NamesDeq;NamesDeq);ac)} )

Lemma: dM-point-subtype
[I,J:fset(ℕ)].  Point(dM(I)) ⊆Point(dM(J)) supposing I ⊆ J

Definition: dM-deq
dM-deq(I) ==  free-dml-deq(names(I);NamesDeq)

Lemma: dM-deq_wf
[I:fset(ℕ)]. (dM-deq(I) ∈ EqDecider(Point(dM(I))))

Definition: dM_inc
<x> ==  <x>

Lemma: dM_inc_wf
[I:fset(ℕ)]. ∀[x:names(I)].  (<x> ∈ Point(dM(I)))

Definition: dM_opp
<1-x> ==  <1-x>

Lemma: dM_opp_wf
[I:fset(ℕ)]. ∀[x:names(I)].  (<1-x> ∈ Point(dM(I)))

Definition: dM0
==  0

Lemma: dM0_wf
[I:fset(ℕ)]. (0 ∈ Point(dM(I)))

Lemma: dM0-sq-empty
[I:Top]. (0 {})

Lemma: dM0-sq-0
[I:Top]. (0 0)

Definition: isdM0
isdM0(x) ==  null(x)

Lemma: isdM0_wf
[I:fset(ℕ)]. ∀[x:Point(dM(I))].  (isdM0(x) ∈ 𝔹)

Lemma: assert-isdM0
[I:fset(ℕ)]. ∀[x:Point(dM(I))].  uiff(↑isdM0(x);x 0 ∈ Point(dM(I)))

Definition: dM1
==  1

Lemma: dM1_wf
[I:fset(ℕ)]. (1 ∈ Point(dM(I)))

Lemma: dM1-sq-singleton-empty
[I:Top]. (1 {{}})

Definition: isdM1
isdM1(x) ==  deq-fset(deq-fset(union-deq(names(I);names(I);NamesDeq;NamesDeq))) {{}}

Lemma: isdM1_wf
[I:fset(ℕ)]. ∀[x:Point(dM(I))].  (isdM1(x) ∈ 𝔹)

Lemma: dM1-meet
[I:fset(ℕ)]. ∀[x:Point(dM(I))].  (1 ∧ x ∈ Point(dM(I)))

Lemma: dM1-meet-dM1
[I:Top]. (1 ∧ 1)

Lemma: assert-isdM1
[I:fset(ℕ)]. ∀[x:Point(dM(I))].  uiff(↑isdM1(x);x 1 ∈ Point(dM(I)))

Lemma: dM_inc_not_1
[I:fset(ℕ)]. ∀[x:names(I)].  (<x> 1 ∈ Point(dM(I))))

Lemma: neg-dM1
[J:fset(ℕ)]. (1) 0 ∈ Point(dM(J)))

Lemma: dma-neg-dM0
[I:Top]. (0) 1)

Lemma: dma-neg-dM1
[I:Top]. (1) 0)

Lemma: dM-0-not-1
[I:fset(ℕ)]. (0 1 ∈ Point(dM(I))))

Lemma: dM1-sq-single-empty
[I:Top]. (1 {{}})

Lemma: dM1-sq-1
[I:Top]. (1 1)

Definition: dMpair
dMpair(i;j) ==  {{inl i,inl j}}

Lemma: dMpair-eq-meet
[I:fset(ℕ)]. ∀[i,j:ℕ].  (dMpair(i;j) = <i> ∧ <j> ∈ Point(dM(I))) supposing (j ∈ and i ∈ I)

Lemma: dMpair_wf
[I:fset(ℕ)]. ∀[i,j:ℕ].  (dMpair(i;j) ∈ Point(dM(I))) supposing (j ∈ and i ∈ I)

Lemma: dM-hom-unique
[I:fset(ℕ)]. ∀[L:BoundedDistributiveLattice]. ∀[eqL:EqDecider(Point(L))]. ∀[g,h:Hom(dM(I);L)].
  h ∈ Hom(dM(I);L) supposing ∀i:names(I). (((g <i>(h <i>) ∈ Point(L)) ∧ ((g <1-i>(h <1-i>) ∈ Point(L)))

Lemma: neg-dM_inc
[I:fset(ℕ)]. ∀[x:names(I)].  (<x>= <1-x> ∈ Point(dM(I)))

Lemma: dma-neg-dM_inc
[I:fset(ℕ)]. ∀[x:names(I)].  (<x>= <1-x> ∈ Point(dM(I)))

Lemma: neg-dM_opp
[I:fset(ℕ)]. ∀[x:names(I)].  (<1-x>= <x> ∈ Point(dM(I)))

Lemma: dma-neg-dM_opp
[I:fset(ℕ)]. ∀[x:names(I)].  (<1-x>= <x> ∈ Point(dM(I)))

Lemma: dM-neg-properties
[I:fset(ℕ)]
  ((∀[x,y:Point(dM(I))].  (x ∧ y) = ¬(x) ∨ ¬(y) ∈ Point(dM(I))))
  ∧ (∀[x,y:Point(dM(I))].  (x ∨ y) = ¬(x) ∧ ¬(y) ∈ Point(dM(I))))
  ∧ (∀[x:Point(dM(I))]. (x)) x ∈ Point(dM(I)))))

Lemma: dma-neg-eq-1-implies-meet-eq-0
[K:fset(ℕ)]. ∀[a2,x1:Point(dM(K))].  ((¬(a2) 1 ∈ Point(dM(K)))  (0 a2 ∧ x1 ∈ Point(dM(K))))

Definition: names-hom
I ⟶ ==  names(J) ⟶ Point(dM(I))

Lemma: names-hom_wf
[I,J:fset(ℕ)].  (I ⟶ J ∈ Type)

Lemma: names-hom-subtype
[I1,J1,I2,J2:fset(ℕ)].  (I1 ⟶ J1 ⊆I2 ⟶ J2) supposing (J2 ⊆ J1 and I1 ⊆ I2)

Definition: dM-lift
dM-lift(I;J;f) ==  free-dma-lift(names(J);NamesDeq;dM(I);free-dml-deq(names(I);NamesDeq);f)

Lemma: dM-lift_wf
[I,J:fset(ℕ)]. ∀[f:I ⟶ J].  (dM-lift(I;J;f) ∈ {g:dma-hom(dM(J);dM(I))| ∀j:names(J). ((g <j>(f j) ∈ Point(dM(I)))} \000C)

Lemma: dM-lift_wf2
[I,J:fset(ℕ)]. ∀[f:I ⟶ J].  (dM-lift(I;J;f) ∈ Point(dM(J)) ⟶ Point(dM(I)))

Lemma: dM-lift-sq
[I',J',I,J,f:Top].  (dM-lift(I;J;f) dM-lift(I';J';f))

Lemma: dM-lift-0-sq
[I,J,f:Top].  (dM-lift(I;J;f) 0)

Lemma: dM-lift-1-sq
[I,J,f:Top].  (dM-lift(I;J;f) 1)

Lemma: dM-lift-0
[I,J:fset(ℕ)]. ∀[f:I ⟶ J].  ((dM-lift(I;J;f) 0) 0 ∈ Point(dM(I)))

Lemma: dM-lift-1
[I,J:fset(ℕ)]. ∀[f:I ⟶ J].  ((dM-lift(I;J;f) 1) 1 ∈ Point(dM(I)))

Lemma: dM-lift-inc
[I,J:fset(ℕ)]. ∀[f:I ⟶ J]. ∀[x:names(J)].  ((dM-lift(I;J;f) <x>(f x) ∈ Point(dM(I)))

Lemma: dM-lift-opp
[I,J:fset(ℕ)]. ∀[f:I ⟶ J]. ∀[x:names(J)].  ((dM-lift(I;J;f) <1-x>= ¬(f x) ∈ Point(dM(I)))

Lemma: dM-lift-meet
[I,J:fset(ℕ)]. ∀[f:I ⟶ J]. ∀[x,y:Point(dM(J))].
  ((dM-lift(I;J;f) x ∧ y) dM-lift(I;J;f) x ∧ dM-lift(I;J;f) y ∈ Point(dM(I)))

Lemma: dM-lift-join
[I,J:fset(ℕ)]. ∀[f:I ⟶ J]. ∀[x,y:Point(dM(J))].
  ((dM-lift(I;J;f) x ∨ y) dM-lift(I;J;f) x ∨ dM-lift(I;J;f) y ∈ Point(dM(I)))

Lemma: dM-lift-neg
[I,J:fset(ℕ)]. ∀[f:I ⟶ J]. ∀[x:Point(dM(J))].  ((dM-lift(I;J;f) ¬(x)) = ¬(dM-lift(I;J;f) x) ∈ Point(dM(I)))

Lemma: dM-lift-unique
[I,J:fset(ℕ)]. ∀[f:I ⟶ J]. ∀[g:dma-hom(dM(J);dM(I))].
  dM-lift(I;J;f) g ∈ dma-hom(dM(J);dM(I)) supposing ∀j:names(J). ((g <j>(f j) ∈ Point(dM(I)))

Lemma: dM-lift-unique-fun
[I,J:fset(ℕ)]. ∀[f:I ⟶ J]. ∀[g:dma-hom(dM(J);dM(I))].
  dM-lift(I;J;f) g ∈ (Point(dM(J)) ⟶ Point(dM(I))) supposing ∀j:names(J). ((g <j>(f j) ∈ Point(dM(I)))

Lemma: dM-lift-dMpair
[I,J:fset(ℕ)]. ∀[f:J ⟶ I]. ∀[x,y:names(I)].  ((dM-lift(J;I;f) dMpair(x;y)) x ∧ y ∈ Point(dM(J)))

Lemma: dM-basis
[I:fset(ℕ)]. ∀[x:Point(dM(I))].  (x \/(λs./\(λx.free-dl-inc(x)"(s))"(x)) ∈ Point(dM(I)))

Lemma: dM-hom-basis
[I:fset(ℕ)]. ∀[x:Point(dM(I))]. ∀[l:BoundedLattice].
  ∀eq:EqDecider(Point(l)). ∀[h:Hom(dM(I);l)]. ((h x) \/(λs./\(λx.(h free-dl-inc(x))"(s))"(x)) ∈ Point(l))

Lemma: dM-homs-equal
[I:fset(ℕ)]. ∀[l:BoundedLattice].
  ∀eq:EqDecider(Point(l))
    ∀[h1,h2:Hom(dM(I);l)].
      h1 h2 ∈ (Point(dM(I)) ⟶ Point(l)) 
      supposing ∀i:names(I). (((h1 <i>(h2 <i>) ∈ Point(l)) ∧ ((h1 <1-i>(h2 <1-i>) ∈ Point(l)))

Lemma: dM-dM-homs-equal
[I,J:fset(ℕ)]. ∀[h1,h2:dma-hom(dM(I);dM(J))].
  h1 h2 ∈ (Point(dM(I)) ⟶ Point(dM(J))) supposing ∀i:names(I). ((h1 <i>(h2 <i>) ∈ Point(dM(J)))

Lemma: dM-hom-invariant
[I:fset(ℕ)]. ∀[J:{J:fset(ℕ)| I ⊆ J} ]. ∀[h:Hom(dM(J);dM(I))].
  ∀[x:Point(dM(I))]. ((h x) x ∈ Point(dM(I))) 
  supposing ∀i:names(I). (((h <i>= <i> ∈ Point(dM(I))) ∧ ((h <1-i>= <1-i> ∈ Point(dM(I))))

Lemma: dM-dma-hom-invariant
[I:fset(ℕ)]. ∀[J:{J:fset(ℕ)| I ⊆ J} ]. ∀[h:dma-hom(dM(J);dM(I))].
  ∀[x:Point(dM(I))]. ((h x) x ∈ Point(dM(I))) supposing ∀i:names(I). ((h <i>= <i> ∈ Point(dM(I)))

Lemma: dM-lift-is-id
[I:fset(ℕ)]. ∀[J:{J:fset(ℕ)| I ⊆ J} ]. ∀[h:I ⟶ J].
  ∀[x:Point(dM(I))]. ((dM-lift(I;J;h) x) x ∈ Point(dM(I))) supposing ∀i:names(I). ((h i) = <i> ∈ Point(dM(I)))

Lemma: dM-subobject
[I,J:fset(ℕ)].  λv.v ∈ dma-hom(dM(I);dM(J)) supposing I ⊆ J

Lemma: dM-lift-is-id2
[I:fset(ℕ)]. ∀[J,K:{K:fset(ℕ)| I ⊆ K} ]. ∀[h:K ⟶ J].
  ∀[x:Point(dM(I))]. ((dM-lift(K;J;h) x) x ∈ Point(dM(K))) supposing ∀i:names(I). ((h i) = <i> ∈ Point(dM(K)))

Lemma: dM-meet-inc-opp
[i:ℕ]. (<i> ∧ <1-i> {{inr ,inl i}})

Lemma: dM-join-inc-opp
[i:ℕ]. (<i> ∨ <1-i> {{inr },{inl i}})

Lemma: one-dimensional-dM
The free DeMorgan algebra with one generator is the free distributive
lattice with two generators, and 1-i. Hence it has exactly six elements.

The proof is somewhat tedious, but we need this lemma so that we
can prove that all paths in "discrete cubical type"  ⌜discr(T)⌝ are constant.⋅

i:ℕ. ∀v:Point(dM({i})).
  ((v 0 ∈ Point(dM({i})))
  ∨ (v 1 ∈ Point(dM({i})))
  ∨ (v = <i> ∈ Point(dM({i})))
  ∨ (v = <1-i> ∈ Point(dM({i})))
  ∨ (v = <i> ∧ <1-i> ∈ Point(dM({i})))
  ∨ (v = <i> ∨ <1-i> ∈ Point(dM({i}))))

Definition: nh-id
==  λx.<x>

Lemma: nh-id_wf
[I:fset(ℕ)]. (1 ∈ I ⟶ I)

Definition: nh-comp
g ⋅ ==  dma-lift-compose(names(I);names(J);NamesDeq;NamesDeq;f;g)

Lemma: nh-comp_wf
[I,J,K:fset(ℕ)]. ∀[f:I ⟶ J]. ∀[g:J ⟶ K].  (g ⋅ f ∈ I ⟶ K)

Lemma: nh-comp-sq
[I,J,f,g:Top].  (g ⋅ dM-lift(I;J;f) g)

Lemma: nh-id-left
I,J:fset(ℕ). ∀f:I ⟶ J.  (f ⋅ f ∈ I ⟶ J)

Lemma: nh-id-right
I,J:fset(ℕ). ∀f:I ⟶ J.  (1 ⋅ f ∈ I ⟶ J)

Lemma: nh-comp-assoc
I,J,K,H:fset(ℕ). ∀f:I ⟶ J. ∀g:J ⟶ K. ∀h:K ⟶ H.  (h ⋅ g ⋅ h ⋅ g ⋅ f ∈ I ⟶ H)

Lemma: nh-comp-cancel
I,J,K:fset(ℕ). ∀f:I ⟶ J. ∀g,h:J ⟶ K. ∀x:names(K).
  (g ⋅ x) (h ⋅ x) ∈ Point(dM(I)) supposing (g x) (h x) ∈ Point(dM(J))

Lemma: nh-comp-is-id
[I,J:fset(ℕ)].
  ∀[f:I ⟶ J]. ∀[g:J ⟶ I].
    g ⋅ 1 ∈ I ⟶ supposing ∀x:names(I). (((g x) = <x> ∈ Point(dM(J))) ∧ ((f x) = <x> ∈ Point(dM(I)))) 
  supposing I ⊆ J

Definition: cube-cat
CubeCat ==  <fset(ℕ), λI,J. I ⟶ J, λI.1, λI,J,K,f,g. g ⋅ f>

Lemma: cube-cat_wf
CubeCat ∈ SmallCategory

Lemma: cube-cat-final
Final({})

Comment: names-cat morphisms doc
=========================== names-cat morphisms ===============================

We define number of special morphisms on the category of names
and prove many equations that they satisfy.

The morphisms we need are
 s, r_i, (i0), (i1), m(i;j), e(i;j), g,i=j, and g,i=1-j.
typical equation 
(that is needed in the proof of the composition operation for the Pi type) is
r_i ⋅ f,i=j f,i=j ⋅ r_j   (see nc-e'-r)⋅

Definition: add-name
I+i ==  fset-add(NamesDeq;i;I)

Lemma: add-name_wf
[I:fset(ℕ)]. ∀[i:ℕ].  (I+i ∈ fset(ℕ))

Lemma: fset-member-add-name
I:fset(ℕ). ∀i,j:ℕ.  uiff(j ∈ I+i;(j i ∈ ℤ) ∨ j ∈ I)

Lemma: trivial-member-add-name1
I:fset(ℕ). ∀i:ℕ.  i ∈ I+i

Lemma: trivial-member-add-name2
I:fset(ℕ). ∀i:ℕ.  ∀[j:ℕ]. i ∈ I+i+j

Lemma: f-subset-add-name1
I:fset(ℕ). ∀[J:fset(ℕ)]. ∀i:ℕ(J ⊆  J ⊆ I+i)

Lemma: f-subset-add-name
I:fset(ℕ). ∀i:ℕ.  I ⊆ I+i

Lemma: add-name-com
[I:fset(ℕ)]. ∀[i,j:ℕ].  (I+i+j I+j+i ∈ fset(ℕ))

Lemma: not-added-name
[I:fset(ℕ)]. ∀[i:ℕ]. ∀[x:names(I+i)].  x ∈ names(I) supposing x ≠ i

Definition: new-name
new-name(I) ==  fset-max(λx.x;I) 1

Lemma: new-name_wf
[I:fset(ℕ)]. (new-name(I) ∈ {i:ℕ| ¬i ∈ I} )

Lemma: new-name-property
[I:fset(ℕ)]. new-name(I) ∈ I)

Lemma: not-new-name
[I:fset(ℕ)]. ∀x:names(I). ((x new-name(I) ∈ ℤ False)

Definition: nc-s
==  λx.<x>

Lemma: nc-s_wf
[I,J:fset(ℕ)].  s ∈ I ⟶ supposing J ⊆ I

Lemma: dM-lift-s
[I,J:fset(ℕ)].  ∀[x:Point(dM(I))]. ((dM-lift(J;I;s) x) x ∈ Point(dM(J))) supposing I ⊆ J

Definition: nc-r
r_i ==  λx.if (x =z i) then <1-i> else <x> fi 

Lemma: nc-r_wf
[I:fset(ℕ)]. ∀[i:ℕ].  r_i ∈ I ⟶ supposing i ∈ I

Definition: nc-p
(i/z) ==  λx.if (x =z i) then else <x> fi 

Lemma: nc-p_wf
[I:fset(ℕ)]. ∀[i:ℕ]. ∀[z:Point(dM(I))].  ((i/z) ∈ I ⟶ I+i)

Lemma: s-comp-nc-p
[I:fset(ℕ)]. ∀[i:ℕ]. ∀[z:Point(dM(I))].  s ⋅ (i/z) 1 ∈ I ⟶ supposing ¬i ∈ I

Definition: nc-0
(i0) ==  λx.if (x =z i) then {} else <x> fi 

Lemma: nc-0_wf
[I:fset(ℕ)]. ∀[i:ℕ].  ((i0) ∈ I ⟶ I+i)

Lemma: nc-0-as-nc-p
[I:fset(ℕ)]. ∀[i:ℕ].  ((i0) (i/0))

Definition: nc-1
(i1) ==  λx.if (x =z i) then {{}} else <x> fi 

Lemma: nc-1_wf
[I:fset(ℕ)]. ∀[i:ℕ].  ((i1) ∈ I ⟶ I+i)

Lemma: nc-1-lemma
J:fset(ℕ). ∀j:ℕ. ∀i:{i:names(J)| ¬i ∈ J} .  (((j1) i) = <i> ∈ Point(dM(J)))

Lemma: nc-1-lemma2
j:ℕ((j1) {{}})

Lemma: dM-lift-nc-1
J:fset(ℕ). ∀j:{i:ℕ| ¬i ∈ J} . ∀v:Point(dM(J)).  ((dM-lift(J;J+j;(j1)) v) v ∈ Point(dM(J)))

Lemma: dM-lift-nc-0
J:fset(ℕ). ∀j:{i:ℕ| ¬i ∈ J} . ∀v:Point(dM(J)).  ((dM-lift(J;J+j;(j0)) v) v ∈ Point(dM(J)))

Lemma: nc-1-as-nc-p
[I:fset(ℕ)]. ∀[i:ℕ].  ((i1) (i/1))

Lemma: r-comp-nc-1
[I:fset(ℕ)]. ∀[i:ℕ].  r_i ⋅ (i1) (i0) ∈ I ⟶ I+i supposing ¬i ∈ I

Lemma: r-comp-nc-0
[I:fset(ℕ)]. ∀[i:ℕ].  r_i ⋅ (i0) (i1) ∈ I ⟶ I+i supposing ¬i ∈ I

Lemma: r-comp-r
[I:fset(ℕ)]. ∀[i:ℕ].  (r_i ⋅ r_i 1 ∈ I+i ⟶ I+i)

Lemma: s-comp-nc-1
[I:fset(ℕ)]. ∀[i:ℕ].  s ⋅ (i1) 1 ∈ I ⟶ supposing ¬i ∈ I

Lemma: s-comp-nc-1-new
[I:fset(ℕ)]. (s ⋅ (new-name(I)1) 1 ∈ I ⟶ I)

Lemma: s-comp-nc-0
[I:fset(ℕ)]. ∀[i:ℕ].  s ⋅ (i0) 1 ∈ I ⟶ supposing ¬i ∈ I

Lemma: s-comp-nc-0-new
[I:fset(ℕ)]. (s ⋅ (new-name(I)0) 1 ∈ I ⟶ I)

Lemma: nc-0-comp-s
[I,K:fset(ℕ)]. ∀[i:ℕ]. ∀[f:K ⟶ I+i].  (i0) ⋅ s ⋅ f ∈ K ⟶ I+i supposing (f i) 0 ∈ Point(dM(K))

Lemma: s-comp-nc-0'
[I:fset(ℕ)]. ∀[i:ℕ]. ∀[j:{j:ℕ| ¬j ∈ I} ].  (s ⋅ (j0) s ∈ I+i ⟶ I)

Lemma: s-comp-nc-1'
[I:fset(ℕ)]. ∀[i:ℕ]. ∀[j:{j:ℕ| ¬j ∈ I} ].  (s ⋅ (j1) s ∈ I+i ⟶ I)

Lemma: s-comp-nc-0-alt
[I:fset(ℕ)]. ∀[i:{i:ℕ| ¬i ∈ I} ]. ∀[j:{j:ℕ| ¬j ∈ I+i} ].  (s ⋅ (i0) s ∈ I+j ⟶ I)

Lemma: s-comp-s
[I:fset(ℕ)]. ∀[i,j:ℕ].  (s ⋅ s ∈ I+i+j ⟶ I)

Definition: nc-m
m(i;j) ==  λx.if (x =z i) then dMpair(i;j) else <x> fi 

Lemma: nc-m_wf
[I:fset(ℕ)]. ∀[i,j:ℕ].  m(i;j) ∈ I+j ⟶ supposing i ∈ I

Lemma: nh-comp-nc-m
[I,K:fset(ℕ)]. ∀[i,j:ℕ]. ∀[f:K ⟶ I+j].  (m(i;j) ⋅ i) i ∧ j ∈ Point(dM(K)) supposing i ∈ I

Lemma: nh-comp-nc-m-eq
[I,K:fset(ℕ)]. ∀[i,j:ℕ]. ∀[f:K ⟶ I+i+j].  (i0) ⋅ s ⋅ s ⋅ m(i;j) ⋅ f ∈ K ⟶ I+i supposing (f i) 0 ∈ Point(dM(K))

Lemma: nh-comp-nc-m-eq2
[I,K:fset(ℕ)]. ∀[i,j:ℕ]. ∀[f:K ⟶ I+i+j].  (i0) ⋅ s ⋅ m(i;j) ⋅ f ∈ K ⟶ I+i supposing (f i) 0 ∈ Point(dM(K))

Lemma: nc-m-nc-1
[I:fset(ℕ)]. ∀[i:{i:ℕ| ¬i ∈ I} ]. ∀[j:{j:ℕ| ¬j ∈ I+i} ].  (m(i;j) ⋅ (j1) 1 ∈ I+i ⟶ I+i)

Lemma: nc-m-nc-0
[I:fset(ℕ)]. ∀[i:{i:ℕ| ¬i ∈ I} ]. ∀[j:{j:ℕ| ¬j ∈ I+i} ].  (m(i;j) ⋅ (j0) (i0) ⋅ s ∈ I+i ⟶ I+i)

Definition: nc-e
e(i;j) ==  λx.if (x =z i) then <j> else <x> fi 

Lemma: nc-e_wf
[I:fset(ℕ)]. ∀[i,j:ℕ].  (e(i;j) ∈ I+j ⟶ I+i)

Lemma: nc-e-same
I:fset(ℕ). ∀z:ℕ.  (e(z;z) 1 ∈ I+z ⟶ I+z)

Lemma: nc-e-comp
I:fset(ℕ). ∀i,j,k:ℕ.  ((¬j ∈ I)  (e(k;j) ⋅ e(j;i) e(k;i) ∈ I+i ⟶ I+k))

Definition: nc-e'
g,i=j ==  λx.if (x =z i) then <j> else fi 

Lemma: nc-e'_wf
[I,J:fset(ℕ)]. ∀[i,j:ℕ]. ∀[g:J ⟶ I].  (g,i=j ∈ J+j ⟶ I+i)

Lemma: nc-e'-lemma1
[I:fset(ℕ)]. ∀[i:ℕ]. ∀[J:fset(ℕ)]. ∀[g:J ⟶ I]. ∀[j:{j:ℕ| ¬j ∈ J} ].  ((i1) ⋅ g,i=j ⋅ (j1) ∈ J ⟶ I+i)

Lemma: nc-e'-lemma2
[I:fset(ℕ)]. ∀[i:ℕ]. ∀[J:fset(ℕ)]. ∀[g:J ⟶ I]. ∀[j:{j:ℕ| ¬j ∈ J} ].  ((i0) ⋅ g,i=j ⋅ (j0) ∈ J ⟶ I+i)

Lemma: nc-e'-lemma3
[I:fset(ℕ)]. ∀[i:{i:ℕ| ¬i ∈ I} ]. ∀[J:fset(ℕ)]. ∀[g:J ⟶ I]. ∀[j:{j:ℕ| ¬j ∈ J} ].  (s ⋅ g,i=j g ⋅ s ∈ J+j ⟶ I)

Lemma: nc-e'-1
[I,i,j:Top].  (1,i=j e(i;j))

Lemma: nc-e-comp-e'
[I,K:fset(ℕ)]. ∀[f:K ⟶ I]. ∀[z,z1,v:ℕ].  f,z=v e(z;z1) ⋅ f,z1=v ∈ K+v ⟶ I+z supposing ¬z1 ∈ I

Lemma: nc-e'-comp-e
[I,J:fset(ℕ)]. ∀[g:J ⟶ I]. ∀[j,k,v:ℕ].  g,j=v ⋅ e(v;k) g,j=k ∈ J+k ⟶ I+j supposing ¬v ∈ J

Lemma: nc-e-comp-nc-p
[I:fset(ℕ)]. ∀[i,j:{j:ℕ| ¬j ∈ I} ]. ∀[z:Point(dM(I))].  (e(i;j) ⋅ (j/z) (i/z) ∈ I ⟶ I+i)

Lemma: nc-e-comp-nc-0
[I:fset(ℕ)]. ∀[i,j:{j:ℕ| ¬j ∈ I} ].  ((i0) e(i;j) ⋅ (j0) ∈ I ⟶ I+i)

Lemma: nc-e-comp-nc-1
[I:fset(ℕ)]. ∀[i,j:{j:ℕ| ¬j ∈ I} ].  ((i1) e(i;j) ⋅ (j1) ∈ I ⟶ I+i)

Lemma: nh-comp-nc-m-eq3
[I:fset(ℕ)]. ∀[i:{i:ℕ| ¬i ∈ I} ]. ∀[j:{j:ℕ| ¬j ∈ I+i} ].  (s,i=j ⋅ (j0) m(i;j) ⋅ (j0) ∈ I+i ⟶ I+i)

Lemma: nh-comp-nc-m-s
[I:fset(ℕ)]. ∀[i:{i:ℕ| ¬i ∈ I} ]. ∀[j:ℕ].  (s ⋅ m(i;j) s ⋅ s ∈ I+i+j ⟶ I)

Lemma: nc-s-comp-e
I:fset(ℕ). ∀i,j:ℕ.  ((¬i ∈ I)  (s ⋅ e(i;j) s ∈ I+j ⟶ I))

Lemma: nc-s-comp-e'
I:fset(ℕ). ∀i,j:ℕ.  ((¬i ∈ I)  (s ⋅ 1,i=j s ∈ I+j ⟶ I))

Lemma: nc-e'-comp-m
[I,J:fset(ℕ)]. ∀[i:{i:ℕ| ¬i ∈ I} ]. ∀[j:{j:ℕ| ¬j ∈ J} ]. ∀[g:J ⟶ I]. ∀[k:{k:ℕ| ¬k ∈ I+i} ]. ∀[l:{l:ℕ| ¬l ∈ J+j} ].
  (g,i=j ⋅ m(j;l) m(i;k) ⋅ g,i=j,k=l ∈ J+j+l ⟶ I+i)

Lemma: nc-e'-comp-m-2
[I,J:fset(ℕ)]. ∀[i:{i:ℕ| ¬i ∈ I} ]. ∀[j:{j:ℕ| ¬j ∈ J} ]. ∀[g:J ⟶ I]. ∀[l:{i:ℕ| ¬i ∈ J+j} ].
  (s ⋅ g,i=j ⋅ m(j;l) g ⋅ s ⋅ s ∈ J+j+l ⟶ I)

Lemma: nc-e'-lemma4
[I,J:fset(ℕ)]. ∀[i:{i:ℕ| ¬i ∈ I} ]. ∀[j:{j:ℕ| ¬j ∈ J} ]. ∀[g:J ⟶ I]. ∀[k:{i1:ℕ| ¬i1 ∈ I+i} ]. ∀[l:{i:ℕ| ¬i ∈ J+j} ].
  ((i0) ⋅ s ⋅ g,i=j,k=l g,i=j ⋅ (j0) ⋅ s ∈ J+j+l ⟶ I+i)

Lemma: nc-e'-lemma5
[I,J:fset(ℕ)]. ∀[i:{i:ℕ| ¬i ∈ I} ]. ∀[j:{j:ℕ| ¬j ∈ J} ]. ∀[g:J ⟶ I]. ∀[k:{i1:ℕ| ¬i1 ∈ I+i} ]. ∀[l:{i:ℕ| ¬i ∈ J+j} ].
  (s ⋅ g,i=j,k=l g ⋅ s ∈ J+j+l ⟶ I)

Lemma: nc-e'-lemma6
[I,J,K:fset(ℕ)]. ∀[f:J ⟶ I]. ∀[g:K ⟶ J]. ∀[z,v,j:ℕ].  f,z=j ⋅ g,j=v f ⋅ g,z=v ∈ K+v ⟶ I+z supposing ¬j ∈ J

Lemma: nc-e'-lemma7
[I:fset(ℕ)]. ∀[i:{i:ℕ| ¬i ∈ I} ].  ∀j:ℕ(1,i=j 1 ⋅ s,i=j ∈ I+j ⟶ I+i)

Definition: nc-r'
g,i=1-j ==  λx.if (x =z i) then <1-j> else fi 

Lemma: nc-r'_wf
[I,J:fset(ℕ)]. ∀[i,j:ℕ]. ∀[g:J ⟶ I].  (g,i=1-j ∈ J+j ⟶ I+i)

Lemma: nc-r'-nc-0
[I:fset(ℕ)]. ∀[i:{i:ℕ| ¬i ∈ I} ]. ∀[J:fset(ℕ)]. ∀[f:J ⟶ I]. ∀[j:{i:ℕ| ¬i ∈ J} ].  ((i1) ⋅ f,i=1-j ⋅ (j0) ∈ J ⟶ I+\000Ci)

Lemma: nc-r'-nc-1
[I:fset(ℕ)]. ∀[i:{i:ℕ| ¬i ∈ I} ]. ∀[J:fset(ℕ)]. ∀[f:J ⟶ I]. ∀[j:{i:ℕ| ¬i ∈ J} ].  ((i0) ⋅ f,i=1-j ⋅ (j1) ∈ J ⟶ I+\000Ci)

Lemma: nc-r'-r
[I:fset(ℕ)]. ∀[i:{i:ℕ| ¬i ∈ I} ]. ∀[J:fset(ℕ)]. ∀[f:J ⟶ I]. ∀[j:{i:ℕ| ¬i ∈ J} ].  (f,i=1-j ⋅ r_j f,i=j ∈ J+j ⟶ I+i)

Lemma: nc-r'-to-e'
[I:fset(ℕ)]. ∀[i:{i:ℕ| ¬i ∈ I} ]. ∀[J:fset(ℕ)]. ∀[f:J ⟶ I]. ∀[j:{i:ℕ| ¬i ∈ J} ].  (f,i=1-j f,i=j ⋅ r_j ∈ J+j ⟶ I+i)

Lemma: nc-e'-r
[I:fset(ℕ)]. ∀[i:{i:ℕ| ¬i ∈ I} ]. ∀[J:fset(ℕ)]. ∀[f:J ⟶ I]. ∀[j:{i:ℕ| ¬i ∈ J} ].
  (r_i ⋅ f,i=j f,i=j ⋅ r_j ∈ J+j ⟶ I+i)

Lemma: nc-r-e'-r
[I:fset(ℕ)]. ∀[i:{i:ℕ| ¬i ∈ I} ]. ∀[J:fset(ℕ)]. ∀[f:J ⟶ I]. ∀[j:{i:ℕ| ¬i ∈ J} ].
  (r_i ⋅ f,i=j ⋅ r_j f,i=j ∈ J+j ⟶ I+i)

Lemma: nc-r'-to-e'2
[I:fset(ℕ)]. ∀[i:{i:ℕ| ¬i ∈ I} ]. ∀[J:fset(ℕ)]. ∀[f:J ⟶ I]. ∀[j:{i:ℕ| ¬i ∈ J} ].  (f,i=1-j r_i ⋅ f,i=j ∈ J+j ⟶ I+i)

Lemma: nc-p-s-commute
[I:fset(ℕ)]. ∀[i,j:ℕ]. ∀[z:Point(dM(I))].  ((i/z) ⋅ s ⋅ (i/z) ∈ I+j ⟶ I+i)

Lemma: nc-1-s-commute
[I:fset(ℕ)]. ∀[i,j:ℕ].  ((i1) ⋅ s ⋅ (i1) ∈ I+j ⟶ I+i)

Lemma: nc-0-s-commute
[I:fset(ℕ)]. ∀[i,j:ℕ].  ((i0) ⋅ s ⋅ (i0) ∈ I+j ⟶ I+i)

Lemma: nc-s-i1-j0
[I:fset(ℕ)]. ∀[i:ℕ]. ∀[j:{j:ℕ| ¬j ∈ I} ].  ((i1) s ⋅ (i1) ⋅ (j0) ∈ I ⟶ I+i)

Lemma: nc-s-i1-j1
[I:fset(ℕ)]. ∀[i:ℕ]. ∀[j:{j:ℕ| ¬j ∈ I} ].  ((i1) s ⋅ (i1) ⋅ (j1) ∈ I ⟶ I+i)

Lemma: nc-e'-p
[I,J:fset(ℕ)]. ∀[f:J ⟶ I]. ∀[z:Point(dM(I))]. ∀[i:ℕ]. ∀[j:{j:ℕ| ¬j ∈ J} ].
  ((i/z) ⋅ f,i=j ⋅ (j/dM-lift(J;I;f) z) ∈ J ⟶ I+i)

Lemma: nc-se'-p
[J:fset(ℕ)]. ∀[k,z:ℕ]. ∀[n:{n:ℕ| ¬n ∈ J} ].  (s,z=n ⋅ (n/<k>e(z;k) ∈ J+k ⟶ J+z)

Lemma: nc-0-s-0
[I:fset(ℕ)]. ∀[i,j:ℕ].  ((i0) ⋅ s ⋅ (i0) s ⋅ (i0) ∈ I+j ⟶ I+i)

Lemma: s-comp-1-e'
[I:fset(ℕ)]. ∀[i:ℕ]. ∀[J:fset(ℕ)]. ∀[f:J ⟶ I]. ∀[j:{j:ℕ| ¬j ∈ I} ]. ∀[k:{j:ℕ| ¬j ∈ J} ].
  (s ⋅ (i1) ⋅ f,j=k (i1) ⋅ f ⋅ s ∈ J+k ⟶ I+i)

Lemma: s-comp-0-e'
[I:fset(ℕ)]. ∀[i:ℕ]. ∀[J:fset(ℕ)]. ∀[f:J ⟶ I]. ∀[j:{j:ℕ| ¬j ∈ I} ]. ∀[k:{j:ℕ| ¬j ∈ J} ].
  (s ⋅ (i0) ⋅ f,j=k (i0) ⋅ f ⋅ s ∈ J+k ⟶ I+i)

Lemma: nc-e'-s-lemma1
[I,J:fset(ℕ)]. ∀[i,z:ℕ]. ∀[g:J ⟶ I]. ∀[j,k:ℕ].  g,i=z ⋅ s ⋅ g,j=k,i=z ∈ J+z+k ⟶ I+i supposing ¬j ∈ I

Lemma: s-comp-if-lemma1
I,J:fset(ℕ). ∀f:J ⟶ I. ∀x:Point(dM(J)).  (s ⋅ λj.if (j =z new-name(I)) then else 1 ⋅ fi  f ∈ J ⟶ I)

Comment: basic cubical sets doc
===========================  basic cubical sets ===============================

This section defines CubicalSet, I-cubes  A(I), and their restrictions
cubical set maps A ⟶ B,  the interval pre-sheaf 𝕀formal-cube(I)
and the maps <rho> and facts about them.

Most of the definitions here are just specializations of definitions
from Error :presheaf models of type theory, where take the category C
in the general definition to be CubeCat. ⋅

Definition: cubical_set
CubicalSet ==  __⊢

Lemma: cubical_set_wf
j⊢ ∈ 𝕌{[i j'']}

Definition: small_cubical_set
SmallCubicalSet ==  small_ps_context{i:l}(CubeCat)

Lemma: small_cubical_set_wf
SmallCubicalSet ∈ 𝕌'

Lemma: small_cubical_set_subtype
small_cubical_set{j:l} ⊆CubicalSet{j}

Lemma: cubical_set_cumulativity
CubicalSet{j} ⊆cubical_set{j':l}

Lemma: cubical_set_cumulativity-i-j
CubicalSet ⊆cubical_set{[j i]:l}

Lemma: cubical_sets_equal
[X,Y:j⊢].
  Y ∈ CubicalSet{j} supposing Y ∈ (F:fset(ℕ) ⟶ 𝕌{j'} × (x:fset(ℕ) ⟶ y:fset(ℕ) ⟶ y ⟶ x ⟶ (F x) ⟶ (F y)))

Definition: ext-eq-cs
X ≡ ==  X ≡ Y

Lemma: ext-eq-cs_wf
[X,Y:j⊢].  (X ≡ Y ∈ ℙ{[i j']})

Definition: I_cube
A(I) ==  I

Lemma: I_cube_wf
[A:j⊢]. ∀[I:fset(ℕ)].  (A(I) ∈ 𝕌{j'})

Lemma: I_cube_pair_redex_lemma
I,G,F:Top.  (<F, G>(I) I)

Definition: cube-set-restriction
f(s) ==  (snd(X)) s

Lemma: cube-set-restriction_wf
[X:j⊢]. ∀[I,J:fset(ℕ)]. ∀[f:J ⟶ I]. ∀[s:X(I)].  (f(s) ∈ X(J))

Lemma: cube_set_restriction_pair_lemma
s,f,J,I,G,F:Top.  (f(s) s)

Lemma: cubical_set-ext
{FM:F:fset(ℕ) ⟶ 𝕌{j'} × (I:fset(ℕ) ⟶ J:fset(ℕ) ⟶ J ⟶ I ⟶ (F I) ⟶ (F J))| 
 let F,M FM 
 in (∀I:fset(ℕ). ∀s:FM(I).  (1(s) s ∈ FM(I)))
    ∧ (∀I,J,K:fset(ℕ). ∀f:J ⟶ I. ∀g:K ⟶ J. ∀s:FM(I).  (f ⋅ g(s) g(f(s)) ∈ FM(K)))}  ≡ CubicalSet{j}

Definition: cs-predicate
cs-predicate(X;I,rho.P[I; rho]) ==  psc-predicate(CubeCat; X; I,rho.P[I; rho])

Lemma: cs-predicate_wf
[X:j⊢]. ∀[P:I:fset(ℕ) ⟶ X(I) ⟶ ℙ{[i j']}].  (cs-predicate(X;I,rho.P[I;rho]) ∈ ℙ{[i j']})

Definition: sub-cubical-set
I,rho.P[I; rho] ==  X|I,rho.P[I; rho]

Lemma: sub-cubical-set_wf
[X:j⊢]. ∀[P:I:fset(ℕ) ⟶ X(I) ⟶ ℙ{j'}].  I,rho.P[I;rho] j⊢ supposing cs-predicate(X;I,rho.P[I;rho])

Lemma: sub-cubical-set_functionality
[X:j⊢]. ∀[P,Q:I:fset(ℕ) ⟶ X(I) ⟶ ℙ].
  I,rho.P[I;rho] ≡ I,rho.Q[I;rho] 
  supposing (∀I:fset(ℕ). ∀rho:X(I).  (P[I;rho] ⇐⇒ Q[I;rho])) ∧ cs-predicate(X;I,rho.P[I;rho])

Lemma: sub-cubical-set-true
[X:j⊢]. I,rho.True ≡ X

Lemma: sub-cubical-set-and
[X:j⊢]. ∀[P,Q:I:fset(ℕ) ⟶ X(I) ⟶ ℙ].
  I,rho.P[I;rho] I,rho.Q[I;rho] ≡ I,rho.P[I;rho] ∧ Q[I;rho] 
  supposing cs-predicate(X;I,rho.P[I;rho]) ∧ cs-predicate(X;I,rho.Q[I;rho])

Definition: cubes
cubes(X) ==  sets(CubeCat; X)

Lemma: cubes_wf
[X:j⊢]. (cubes(X) ∈ small-category{[i j']:l})

Lemma: cubes-ob
[X:Top]. (cat-ob(cubes(X)) I:fset(ℕ) × X(I))

Lemma: cubes-id
[X:Top]. (cat-id(cubes(X)) ~ λ_.1)

Lemma: cubes-arrow
[X:Top]. (cat-arrow(cubes(X)) ~ λAa,Bb. let A,a Aa in let B,b Bb in {f:A ⟶ B| f(b) a ∈ (X A)} )

Lemma: cubes-comp
[X:Top]. (cat-comp(cubes(X)) ~ λAa,Bb,_,f,g. g ⋅ f)

Definition: interval-presheaf
This is the cubical set that assigns to finite set of names the 
points of the free DeMorgan algebra generated by I.
Since, morphism I->maps -> dM(I), it lifts to 
homomorphism dM(J)->dM(I).⋅

𝕀 ==  <λI.Point(dM(I)), λJ,I,f. dM-lift(I;J;f)>

Lemma: interval-presheaf_wf
𝕀 ∈ SmallCubicalSet

Lemma: interval-presheaf-restriction
[I,K,f,r:Top].  (f(r) dM-lift(K;I;f) r)

Lemma: nc-e'-p2
[I,J:fset(ℕ)]. ∀[f:J ⟶ I]. ∀[r:Point(dM(I))]. ∀[i:ℕ]. ∀[j:{j:ℕ| ¬j ∈ J} ].  (f,i=j ⋅ (j/f(r)) (i/r) ⋅ f ∈ J ⟶ I+i)

Definition: formal-cube
formal-cube(I) ==  <λJ.J ⟶ I, λI,J,f,g. g ⋅ f>

Lemma: formal-cube-is-rep-pre-sheaf
[I:Top]. (formal-cube(I) rep-pre-sheaf(CubeCat;I))

Lemma: formal-cube_wf
[I:fset(ℕ)]. formal-cube(I) ⊢

Lemma: formal-cube_wf1
[I:fset(ℕ)]. formal-cube(I) j⊢

Lemma: formal-cube-restriction
[I,f,s,A,B:Top].  (f(s) s ⋅ f)

Lemma: formal-cube-singleton1
x:ℕA,f. (f x) ∈ nat-trans(op-cat(CubeCat);TypeCat';formal-cube({x});𝕀))

Lemma: formal-cube-singleton2
x:ℕA,u,x. u ∈ nat-trans(op-cat(CubeCat);TypeCat';𝕀;formal-cube({x})))

Lemma: formal-cube-singleton-iso-interval-presheaf
x:ℕcat-isomorphic(FUN(op-cat(CubeCat);TypeCat');formal-cube({x});𝕀)

Definition: representable-cube-set
representable-cube-set() ==  {G:cubical_set{1:l}| ∃I:fset(ℕ). (G formal-cube(I) ∈ cubical_set{1:l})} 

Lemma: representable-cube-set_wf
representable-cube-set() ∈ 𝕌{3}

Definition: trivial-cube-set
() ==  <λJ.Unit, λJ,K,f,x. ⋅>

Lemma: trivial-cube-set_wf
() j⊢

Definition: discrete-cube
discrete-cube(A) ==  <λJ.A, λJ,K,f,x. x>

Lemma: discrete-cube_wf
[A:𝕌{j}]. discrete-cube(A) j⊢

Definition: cubical-false
False ==  discrete-cube(Void)

Lemma: cubical-false_wf
False j⊢

Definition: cube_set_map
A ⟶ ==  A ⟶ B

Lemma: cube_set_map_wf
[A,B:j⊢].  (A j⟶ B ∈ 𝕌{j''})

Lemma: cube_set_map_cumulativity
[G,H:j⊢].  (H j⟶ G ⊆cube_set_map{j':l}(H; G))

Lemma: cube_set_map_cumulativity-i-j
[G,H:⊢].  (H ⟶ G ⊆ij⟶ G)

Definition: context-map
<rho> ==  λA,f. (Gamma rho)

Lemma: context-map_wf
I:fset(ℕ). ∀Gamma:j⊢. ∀rho:Gamma(I).  (<rho> ∈ formal-cube(I) j⟶ Gamma)

Lemma: csm-equal
[A,B:j⊢]. ∀[f:A j⟶ B]. ∀[g:I:fset(ℕ) ⟶ A(I) ⟶ B(I)].  g ∈ j⟶ supposing g ∈ (I:fset(ℕ) ⟶ A(I) ⟶ B(I))

Lemma: csm-equal2
[A,B:j⊢]. ∀[f,g:A j⟶ B].  g ∈ j⟶ supposing ∀K:fset(ℕ). ∀x:A(K).  ((f x) (g x) ∈ B(K))

Lemma: cube-set-map-subtype
[A,B:j⊢].  (A j⟶ B ⊆(I:fset(ℕ) ⟶ A(I) ⟶ B(I)))

Definition: csm-ap
(s)x ==  x

Lemma: csm-ap_wf
[X,Y:j⊢]. ∀[s:X j⟶ Y]. ∀[I:fset(ℕ)]. ∀[x:X(I)].  ((s)x ∈ Y(I))

Lemma: csm-ap-context-map
[I,J,a,rho,G:Top].  ((<rho>)a a(rho))

Definition: csm-comp
==  λx.((G x) (F x))

Lemma: csm-comp-sq
[A,B,C,F,G:Top].  (G G)

Lemma: csm-comp_wf
[A,B,C:j⊢]. ∀[F:A j⟶ B]. ∀[G:B j⟶ C].  (G F ∈ j⟶ C)

Lemma: csm-comp-assoc
[A,B,C,D:j⊢]. ∀[F:A j⟶ B]. ∀[G:B j⟶ C]. ∀[H:C j⟶ D].  (H F ∈ j⟶ D)

Definition: csm-id
1(X) ==  λA,x. x

Lemma: csm-id-sq
[X:j⊢]. (1(X) identity-trans(op-cat(CubeCat);type-cat{j':l};X))

Lemma: csm-id_wf
[X:j⊢]. (1(X) ∈ j⟶ X)

Lemma: csm-id-comp
[A,B:j⊢]. ∀[sigma:A j⟶ B].  ((sigma 1(A) sigma ∈ j⟶ B) ∧ (1(B) sigma sigma ∈ j⟶ B))

Lemma: csm-id-comp-sq
[A,B,C,D,x,y,X,Y,Z:Top].  (1(D) y)

Lemma: csm-ap-csm-comp
Gamma,Delta,Z,s1,s2,I,alpha:Top.  ((s2 s1)alpha (s2)(s1)alpha)

Lemma: context-map-1
[I:fset(ℕ)]. (<1> 1(formal-cube(I)) ∈ formal-cube(I) ⟶ formal-cube(I))

Lemma: context-map-comp
[I,J,K:fset(ℕ)]. ∀[f:J ⟶ I]. ∀[g:K ⟶ J].  (<f ⋅ g> = <f> o <g> ∈ formal-cube(K) ⟶ formal-cube(I))

Lemma: cube-set-restriction-id
[X:j⊢]. ∀[I:fset(ℕ)]. ∀[s:X(I)].  (1(s) s ∈ X(I))

Lemma: cube-set-restriction-when-id
[X:j⊢]. ∀[I:fset(ℕ)]. ∀[s:X(I)]. ∀[f:I ⟶ I].  f(s) s ∈ X(I) supposing 1 ∈ I ⟶ I

Lemma: cube-set-restriction-comp
X:j⊢. ∀I,J,K:fset(ℕ). ∀f:J ⟶ I. ∀g:K ⟶ J. ∀a:X(I).  (g(f(a)) f ⋅ g(a) ∈ X(K))

Lemma: cube-set-restriction-comp2
X:j⊢. ∀I,J1,J2,K:fset(ℕ). ∀f:J1 ⟶ I. ∀g:K ⟶ J1. ∀a:X(I).  g(f(a)) f ⋅ g(a) ∈ X(K) supposing J1 J2 ∈ fset(ℕ)

Lemma: cube-set-map-is
[A,B:j⊢].
  (A j⟶ {trans:I:fset(ℕ) ⟶ A(I) ⟶ B(I)| 
              ∀I,J:fset(ℕ). ∀g:J ⟶ I.  ((λs.g(trans s)) s.(trans g(s))) ∈ (A(I) ⟶ B(J)))} )

Lemma: csm-ap-restriction
X,Y:j⊢. ∀s:X j⟶ Y. ∀I,J:fset(ℕ). ∀f:J ⟶ I. ∀a:X(I).  (f((s)a) (s)f(a) ∈ Y(J))

Definition: sub_cubical_set
Y ⊆ ==  1(Y) ∈ Y ⟶ X

Lemma: sub_cubical_set_wf
[Y,X:j⊢].  (sub_cubical_set{j:l}(Y; X) ∈ 𝕌{[i j'']})

Lemma: sub_cubical_set-cumulativity1
[Y,X:j⊢].  sub_cubical_set{[j i]:l}(Y; X) supposing sub_cubical_set{j:l}(Y; X)

Lemma: sub_cubical_set_transitivity
[X,Y,Z:j⊢].  sub_cubical_set{j:l}(Z; X) supposing sub_cubical_set{j:l}(Z; Y) ∧ sub_cubical_set{j:l}(Y; X)

Lemma: sub_cubical_set_weakening
[X,Z:j⊢].  sub_cubical_set{j:l}(Z; X) supposing X ∈ CubicalSet{j}

Lemma: sub_cubical_set_self
[X:j⊢]. sub_cubical_set{j:l}(X; X)

Lemma: implies-sub_cubical_set
[Y,X:j⊢].
  (sub_cubical_set{j:l}(Y; X)) supposing 
     ((∀A,B:fset(ℕ). ∀g:B ⟶ A. ∀rho:Y(A).  (g(rho) g(rho) ∈ X(B))) and 
     (∀I:fset(ℕ). (Y(I) ⊆X(I))))

Lemma: subset-I_cube
[X,Y:j⊢].  ∀I:fset(ℕ). (Y(I) ⊆X(I)) supposing sub_cubical_set{j:l}(Y; X)

Lemma: subset-restriction
[X,Y:j⊢].  ∀[I,J:fset(ℕ)]. ∀[f:J ⟶ I]. ∀[a:Y(I)].  (f(a) f(a) ∈ X(J)) supposing sub_cubical_set{j:l}(Y; X)

Lemma: cube_set_map_subtype
[X,Y,Z:j⊢].  j⟶ X ⊆j⟶ supposing sub_cubical_set{j:l}(X; Y)

Lemma: cube_set_map_subtype2
[X,Y,Z:j⊢].  j⟶ Z ⊆j⟶ supposing sub_cubical_set{j:l}(Y; X)

Lemma: cube_set_map_subtype3
[X,Y,Z,U:j⊢].  (X j⟶ Z ⊆j⟶ U) supposing (sub_cubical_set{j:l}(Y; X) and sub_cubical_set{j:l}(Z; U))

Lemma: cube_set_map_cumulativity2
[X,Y,Z,U:j⊢].  (X j⟶ Z ⊆j⟶ U) supposing (sub_cubical_set{j:l}(Y; X) and sub_cubical_set{j:l}(Z; U))

Comment: face lattice doc
=================================  face lattice ================================

The face lattice for names is
 ⌜face-lattice(names(I);NamesDeq)⌝ then free distributive lattice
on generators ⌜(i=0)⌝ and ⌜(i=1)⌝ with ⌜(i=0) ∧ (i=1) 0 ∈ Point(face-lattice(names(I);NamesDeq))⌝⋅

Definition: face_lattice
face_lattice(I) ==  face-lattice(names(I);NamesDeq)

Lemma: face_lattice_wf
[I:fset(ℕ)]. (face_lattice(I) ∈ BoundedDistributiveLattice)

Lemma: face_lattice-point_wf
[I:fset(ℕ)]. (Point(face_lattice(I)) ∈ Type)

Lemma: respects-equality-face-lattice-point
[I,J:fset(ℕ)].  respects-equality(Point(face_lattice(I));Point(face_lattice(J)))

Lemma: respects-equality-face-lattice-point-2
[I,J:fset(ℕ)].  respects-equality(fset(fset(names(I) names(I)));Point(face_lattice(J)))

Definition: face_lattice-deq
face_lattice-deq() ==  deq-fset(deq-fset(sumdeq(NamesDeq;NamesDeq)))

Lemma: face_lattice-deq_wf
[I:fset(ℕ)]. (face_lattice-deq() ∈ EqDecider(Point(face_lattice(I))))

Lemma: decidable__equal_face_lattice
[I:fset(ℕ)]. ∀x,y:Point(face_lattice(I)).  Dec(x y ∈ Point(face_lattice(I)))

Lemma: face-lattice-0-not-1
[J:fset(ℕ)]. (0 1 ∈ Point(face_lattice(J))))

Lemma: face_lattice-1-join-irreducible
I:fset(ℕ). ∀x,y:Point(face_lattice(I)).
  (x ∨ 1 ∈ Point(face_lattice(I)) ⇐⇒ (x 1 ∈ Point(face_lattice(I))) ∨ (y 1 ∈ Point(face_lattice(I))))

Lemma: face_lattice-fset-join-eq-1
I:fset(ℕ). ∀s:fset(Point(face_lattice(I))).
  (\/(s) 1 ∈ Point(face_lattice(I)) ⇐⇒ ∃x:Point(face_lattice(I)). (x ∈ s ∧ (x 1 ∈ Point(face_lattice(I)))))

Lemma: face_lattice-point-subtype
[I,J:fset(ℕ)].  Point(face_lattice(I)) ⊆Point(face_lattice(J)) supposing I ⊆ J

Lemma: face_lattice-join-invariant
[v,y,I,J:Top].  (v ∨ v ∨ y)

Lemma: face_lattice-meet-invariant
[v,y,I,J:Top].  (v ∧ v ∧ y)

Definition: fl0
(x=0) ==  (x=0)

Lemma: fl0_wf
[I:fset(ℕ)]. ∀[x:names(I)].  ((x=0) ∈ Point(face_lattice(I)))

Definition: fl1
(x=1) ==  (x=1)

Lemma: fl1_wf
[I:fset(ℕ)]. ∀[x:names(I)].  ((x=1) ∈ Point(face_lattice(I)))

Definition: fl-eq
(x==y) ==  free-dml-deq(names(I);NamesDeq) y

Lemma: fl-eq_wf
[I:fset(ℕ)]. ∀[x,y:Point(face_lattice(I))].  ((x==y) ∈ 𝔹)

Lemma: fl-eq-0-1-false
I:Top. ((0==1) ff)

Lemma: assert-fl-eq
[I:fset(ℕ)]. ∀[x,y:Point(face_lattice(I))].  uiff(↑(x==y);x y ∈ Point(face_lattice(I)))

Definition: fl-deq
Deq(F(I)) ==  λx,y. (x==y)

Lemma: face_lattice-induction
I:fset(ℕ)
  ∀[P:Point(face_lattice(I)) ⟶ ℙ]
    ((∀x:Point(face_lattice(I)). SqStable(P[x]))
     P[0]
     P[1]
     (∀x,y:Point(face_lattice(I)).  (P[x]  P[y]  P[x ∨ y]))
     (∀x:Point(face_lattice(I)). (P[x]  (∀i:names(I). (P[(i=0) ∧ x] ∧ P[(i=1) ∧ x]))))
     {∀x:Point(face_lattice(I)). P[x]})

Lemma: face_lattice-hom-fixes-sublattice
[I,J:fset(ℕ)].
  ∀[f:Hom(face_lattice(J);face_lattice(I))]. ∀[x:Point(face_lattice(J))].
    (f x) x ∈ Point(face_lattice(I)) 
    supposing ∀i:names(J)
                (((f (i=0)) (i=0) ∈ Point(face_lattice(I))) ∧ ((f (i=1)) (i=1) ∈ Point(face_lattice(I)))) 
  supposing J ⊆ I

Lemma: face_lattice-hom-fixes-sublattice2
[I,J:fset(ℕ)].
  ∀[f:Hom(face_lattice(I);face_lattice(J))]. ∀[x:Point(face_lattice(J))].
    (f x) x ∈ Point(face_lattice(J)) 
    supposing ∀i:names(J)
                (((f (i=0)) (i=0) ∈ Point(face_lattice(J))) ∧ ((f (i=1)) (i=1) ∈ Point(face_lattice(J)))) 
  supposing J ⊆ I

Lemma: fl-deq_wf
[I:fset(ℕ)]. (Deq(F(I)) ∈ EqDecider(Point(face_lattice(I))))

Lemma: fl0-not-1
[I:fset(ℕ)]. ∀[x:names(I)].  ((x=0) 1 ∈ Point(face_lattice(I))))

Lemma: fl1-not-1
[I:fset(ℕ)]. ∀[x:names(I)].  ((x=1) 1 ∈ Point(face_lattice(I))))

Lemma: FL-meet-0-1
[I:fset(ℕ)]. ∀[x:names(I)].
  (((x=0) ∧ (x=1) 0 ∈ Point(face_lattice(I))) ∧ ((x=1) ∧ (x=0) 0 ∈ Point(face_lattice(I))))

Lemma: face_lattice-hom-equal
[I,J:fset(ℕ)]. ∀[g,h:Hom(face_lattice(I);face_lattice(J))].
  h ∈ Hom(face_lattice(I);face_lattice(J)) 
  supposing (∀x:names(I). ((g (x=0)) (h (x=0)) ∈ Point(face_lattice(J))))
  ∧ (∀x:names(I). ((g (x=1)) (h (x=1)) ∈ Point(face_lattice(J))))

Definition: irr_face
irr_face(I;as;bs) ==  /\(λa.(a=0)"(as) ⋃ λb.(b=1)"(bs))

Lemma: irr_face_wf
[I:fset(ℕ)]. ∀[as,bs:fset(names(I))].  (irr_face(I;as;bs) ∈ Point(face_lattice(I)))

Definition: face_lattice_components
face_lattice_components(I;x) ==  λs.<fset-mapfilter(λx.outl(x);λx.isl(x);s), fset-mapfilter(λx.outr(x);λx.isr(x);s)>"(x)

Lemma: face_lattice_components_wf
I:fset(ℕ). ∀x:Point(face_lattice(I)).
  (face_lattice_components(I;x) ∈ {fs:fset({p:fset(names(I)) × fset(names(I))| ↑fset-disjoint(NamesDeq;fst(p);snd(p))} )\000C| 
                                   \/(λpr.irr_face(I;fst(pr);snd(pr))"(fs)) ∈ Point(face_lattice(I))} )

Lemma: face_lattice-basis
I:fset(ℕ). ∀x:Point(face_lattice(I)).
  ∃fs:fset({p:fset(names(I)) × fset(names(I))| ↑fset-disjoint(NamesDeq;fst(p);snd(p))} )
   (x \/(λpr.irr_face(I;fst(pr);snd(pr))"(fs)) ∈ Point(face_lattice(I)))

Definition: dM-to-FL
dM-to-FL(I;z) ==
  lattice-extend(face_lattice(I);union-deq(names(I);names(I);NamesDeq;NamesDeq);face_lattice-deq();λx.case x
                                                                                                       of inl(a) =>
                                                                                                       (a=1)
                                                                                                       inr(a) =>
                                                                                                       (a=0);z)

Lemma: dM-to-FL_wf
[I:fset(ℕ)]. ∀[z:Point(dM(I))].  (dM-to-FL(I;z) ∈ Point(face_lattice(I)))

Lemma: dM-to-FL-sq
[I,J,v:Top].  (dM-to-FL(I;v) dM-to-FL(J;v))

Lemma: dm-neg-sq
[I,J,v:Top].  (v) ~ ¬(v))

Lemma: dM-to-FL-inc
[I:fset(ℕ)]. ∀[x:names(I)].  (dM-to-FL(I;<x>(x=1) ∈ Point(face_lattice(I)))

Lemma: dM-to-FL-opp
[I:fset(ℕ)]. ∀[x:names(I)].  (dM-to-FL(I;<1-x>(x=0) ∈ Point(face_lattice(I)))

Lemma: dM-to-FL-is-hom
[I:fset(ℕ)]. z.dM-to-FL(I;z) ∈ Hom(free-DeMorgan-lattice(names(I);NamesDeq);face_lattice(I)))

Lemma: dM-to-FL-unique
[I:fset(ℕ)]. ∀[g:Hom(free-DeMorgan-lattice(names(I);NamesDeq);face_lattice(I))].
  ∀[x:Point(dM(I))]. (dM-to-FL(I;x) (g x) ∈ Point(face_lattice(I))) 
  supposing ∀i:names(I). (((g <i>(i=1) ∈ Point(face_lattice(I))) ∧ ((g <1-i>(i=0) ∈ Point(face_lattice(I))))

Lemma: dM-to-FL-properties
[I:fset(ℕ)]
  ((∀x,y:Point(free-DeMorgan-lattice(names(I);NamesDeq)).
      (dM-to-FL(I;x ∨ y) dM-to-FL(I;x) ∨ dM-to-FL(I;y) ∈ Point(face_lattice(I))))
  ∧ (∀x,y:Point(free-DeMorgan-lattice(names(I);NamesDeq)).
       (dM-to-FL(I;x ∧ y) dM-to-FL(I;x) ∧ dM-to-FL(I;y) ∈ Point(face_lattice(I))))
  ∧ (dM-to-FL(I;0) 0 ∈ Point(face_lattice(I)))
  ∧ (dM-to-FL(I;1) 1 ∈ Point(face_lattice(I))))

Lemma: dM-to-FL-eq-1
[I:fset(ℕ)]. ∀[x:Point(dM(I))].  uiff(dM-to-FL(I;x) 1 ∈ Point(face_lattice(I));x 1 ∈ Point(dM(I)))

Lemma: dM-to-FL-1
[I:fset(ℕ)]. (dM-to-FL(I;1) 1 ∈ Point(face_lattice(I)))

Lemma: dM-to-FL-dM1
[I:fset(ℕ)]. (dM-to-FL(I;1) 1 ∈ Point(face_lattice(I)))

Lemma: dM-to-FL-dM0
[I:fset(ℕ)]. (dM-to-FL(I;0) 0 ∈ Point(face_lattice(I)))

Lemma: dM-to-FL-neg
[I:fset(ℕ)]
  ∀x:Point(free-DeMorgan-lattice(names(I);NamesDeq)). (dM-to-FL(I;x) ∧ dM-to-FL(I;¬(x)) 0 ∈ Point(face_lattice(I)))

Lemma: dM-to-FL-neg2
[I:fset(ℕ)]
  ∀x:Point(free-DeMorgan-lattice(names(I);NamesDeq)). (dM-to-FL(I;¬(x)) ∧ dM-to-FL(I;x) 0 ∈ Point(face_lattice(I)))

Lemma: dM-to-FL-neg-is-1
[I:fset(ℕ)]
  ∀x:Point(free-DeMorgan-lattice(names(I);NamesDeq))
    dM-to-FL(I;x) 0 ∈ Point(face_lattice(I)) supposing dM-to-FL(I;¬(x)) 1 ∈ Point(face_lattice(I))

Lemma: face-lattice-hom-is-id
I:fset(ℕ)
  ∀[h:Hom(face_lattice(I);face_lattice(I))]
    x.x) ∈ Hom(face_lattice(I);face_lattice(I)) 
    supposing ∀x:names(I). (((h (x=0)) (x=0) ∈ Point(face_lattice(I))) ∧ ((h (x=1)) (x=1) ∈ Point(face_lattice(I))))

Definition: fl-morph
<f> ==  fl-lift(names(J);NamesDeq;face_lattice(I);face_lattice-deq();λj.dM-to-FL(I;¬(f j));λj.dM-to-FL(I;f j))

Lemma: fl-morph_wf
[I,J:fset(ℕ)]. ∀[f:I ⟶ J].  (<f> ∈ Hom(face_lattice(J);face_lattice(I)))

Lemma: fl-morph-id
I:fset(ℕ). (<1> x.x) ∈ (Point(face_lattice(I)) ⟶ Point(face_lattice(I))))

Lemma: apply-fl-morph-id
[I:fset(ℕ)]. ∀[phi:Point(face_lattice(I))].  ((phi)<1> phi ∈ Point(face_lattice(I)))

Lemma: fl-morph-face-lattice0
[J,I:fset(ℕ)]. ∀[f:J ⟶ I]. ∀[x:names(I)].  (((x=0))<f> dM-to-FL(J;¬(f x)) ∈ Point(face_lattice(J)))

Lemma: fl-morph-face-lattice1
[J,I:fset(ℕ)]. ∀[f:J ⟶ I]. ∀[x:names(I)].  (((x=1))<f> dM-to-FL(J;f x) ∈ Point(face_lattice(J)))

Lemma: fl-morph-fl0
[J,I:fset(ℕ)]. ∀[f:J ⟶ I]. ∀[x:names(I)].  (((x=0))<f> dM-to-FL(J;¬(f x)) ∈ Point(face_lattice(J)))

Lemma: fl-morph-fl0-is-1
[J,I:fset(ℕ)]. ∀[f:J ⟶ I]. ∀[x:names(I)].  uiff(((x=0))<f> 1 ∈ Point(face_lattice(J));(f x) 0 ∈ Point(dM(J)))

Lemma: fl-s-fl0
[I:fset(ℕ)]. ∀[x:names(I)].  (((x=0))<s> (x=0) ∈ Point(face_lattice(I+x)))

Lemma: fl-morph-fl1
[J,I:fset(ℕ)]. ∀[f:J ⟶ I]. ∀[x:names(I)].  (((x=1))<f> dM-to-FL(J;f x) ∈ Point(face_lattice(J)))

Lemma: fl-morph-fl1-is-1
[J,I:fset(ℕ)]. ∀[f:J ⟶ I]. ∀[x:names(I)].  uiff(((x=1))<f> 1 ∈ Point(face_lattice(J));(f x) 1 ∈ Point(dM(J)))

Lemma: fl-s-fl1
[I:fset(ℕ)]. ∀[x:names(I)].  (((x=1))<s> (x=1) ∈ Point(face_lattice(I+x)))

Lemma: fl-morph-comp-1
[J,K:fset(ℕ)]. ∀[f:K ⟶ J]. ∀[z:Point(dM(J))]. ∀[h:dma-hom(dM(J);dM(K))].
  (dM-to-FL(J;z))<f> dM-to-FL(K;h z) ∈ Point(face_lattice(K)) supposing ∀i:names(J). ((h <i>(f i) ∈ Point(dM(K)))

Lemma: fl-morph-comp-dM-lift
[I,J:fset(ℕ)]. ∀[f:J ⟶ I]. ∀[z:Point(dM(I))].
  ((dM-to-FL(I;z))<f> dM-to-FL(J;dM-lift(J;I;f) z) ∈ Point(face_lattice(J)))

Lemma: fl-morph-comp
[I,J,K:fset(ℕ)]. ∀[f:J ⟶ I]. ∀[g:K ⟶ J].
  (<f ⋅ g> (<g> o <f>) ∈ (Point(face_lattice(I)) ⟶ Point(face_lattice(K))))

Lemma: fl-morph-comp2
[I,J,K:fset(ℕ)]. ∀[f:J ⟶ I]. ∀[g:K ⟶ J]. ∀[x:Point(face_lattice(I))].
  (((x)<f>)<g> (x)<f ⋅ g> ∈ Point(face_lattice(K)))

Lemma: fl-morph-1
[A,B:fset(ℕ)]. ∀[g:A ⟶ B].  ((1)<g> 1 ∈ Point(face_lattice(A)))

Lemma: fl-morph-0
[A,B:fset(ℕ)]. ∀[g:A ⟶ B].  ((0)<g> 0 ∈ Point(face_lattice(A)))

Lemma: fl-morph-join
[A,B:fset(ℕ)]. ∀[g:A ⟶ B]. ∀[x,y:Point(face_lattice(B))].  ((x ∨ y)<g> (x)<g> ∨ (y)<g> ∈ Point(face_lattice(A)))

Lemma: fl-morph-meet
[A,B:fset(ℕ)]. ∀[g:A ⟶ B]. ∀[x,y:Point(face_lattice(B))].  ((x ∧ y)<g> (x)<g> ∧ (y)<g> ∈ Point(face_lattice(A)))

Lemma: fl-morph-fset-meet
[A,B:fset(ℕ)]. ∀[g:A ⟶ B]. ∀[x:fset(Point(face_lattice(B)))].  ((/\(x))<g> /\(<g>"(x)) ∈ Point(face_lattice(A)))

Lemma: fl-morph-fset-join
[A,B:fset(ℕ)]. ∀[g:A ⟶ B]. ∀[x:fset(Point(face_lattice(B)))].  ((\/(x))<g> \/(<g>"(x)) ∈ Point(face_lattice(A)))

Definition: irr-face-morph
irr-face-morph(I;as;bs) ==  λi.if i ∈b as then if i ∈b bs then else <i> fi 

Lemma: irr-face-morph_wf
[I:fset(ℕ)]. ∀[as,bs:fset(names(I))].  (irr-face-morph(I;as;bs) ∈ I ⟶ I)

Definition: face-presheaf
𝔽 ==  <λI.Point(face_lattice(I)), λJ,I,f. <f>>

Lemma: face-presheaf_wf1
𝔽 ∈ SmallCubicalSet

Lemma: face-presheaf_wf
𝔽 ∈ small_cubical_set{j:l}

Lemma: face-presheaf_wf2
𝔽 j⊢

Lemma: face-presheaf-point-subtype
[I:fset(ℕ)]. (𝔽(I) ⊆Point(face_lattice(I)))

Lemma: subtype-face-presheaf-point
[I:fset(ℕ)]. (Point(face_lattice(I)) ⊆r 𝔽(I))

Lemma: face-fl-morph-id
[I:fset(ℕ)]. ∀[phi:𝔽(I)].  ((phi)<1> phi ∈ 𝔽(I))

Lemma: face-presheaf-restriction-1
[A,B:fset(ℕ)]. ∀[g:B ⟶ A].  (g(1) 1 ∈ Point(face_lattice(B)))

Lemma: fl-morph-restriction
[I,J,K:fset(ℕ)]. ∀[f:I ⟶ J]. ∀[g:J ⟶ K]. ∀[phi:𝔽(K)].  ((g(phi))<f> g ⋅ f(phi) ∈ 𝔽(I))

Definition: fl-join
fl-join(I;x;y) ==  x ∨ y

Lemma: fl-join_wf
[I:fset(ℕ)]. ∀[x,y:𝔽(I)].  (fl-join(I;x;y) ∈ 𝔽(I))

Definition: dM_FL
dM_FL() ==  λI,z. dM-to-FL(I;z)

Lemma: dM_FL_wf
dM_FL() ∈ 𝕀 j⟶ 𝔽

Lemma: face_lattice_hom_subtype
[I,J:fset(ℕ)].  Hom(face_lattice(J);face_lattice(I)) ⊆Hom(face_lattice(I);face_lattice(I)) supposing I ⊆ J

Definition: fl-all-hom
fl-all-hom(I;i) ==
  fl-lift(names(I+i);NamesDeq;face_lattice(I);face_lattice-deq();λx.if (x =z i) then else (x=0) fi x.if (x =z i)
                                                                                                         then 0
                                                                                                         else (x=1)
                                                                                                         fi )

Lemma: fl-all-hom_wf1
[I:fset(ℕ)]. ∀[i:ℕ].
  (fl-all-hom(I;i) ∈ {g:Hom(face_lattice(I+i);face_lattice(I))| 
                      (∀j:names(I)
                         ((¬(j i ∈ ℤ))
                          (((g (j=0)) (j=0) ∈ Point(face_lattice(I)))
                            ∧ ((g (j=1)) (j=1) ∈ Point(face_lattice(I))))))
                      ∧ ((g (i=0)) 0 ∈ Point(face_lattice(I)))
                      ∧ ((g (i=1)) 0 ∈ Point(face_lattice(I)))} )

Lemma: fl-all-hom_wf
[I:fset(ℕ)]. ∀[i:{i:ℕ| ¬i ∈ I} ].
  (fl-all-hom(I;i) ∈ {g:Hom(face_lattice(I+i);face_lattice(I))| 
                      (∀x:Point(face_lattice(I)). ((g x) x ∈ Point(face_lattice(I))))
                      ∧ ((g (i=0)) 0 ∈ Point(face_lattice(I)))
                      ∧ ((g (i=1)) 0 ∈ Point(face_lattice(I)))} )

Definition: fl_all
(∀i.phi) ==  fl-all-hom(I;i) phi

Lemma: fl_all_wf
[I:fset(ℕ)]. ∀[i:ℕ]. ∀[phi:Point(face_lattice(I+i))].  ((∀i.phi) ∈ Point(face_lattice(I)))

Lemma: fl_all-1
I,J,i:Top.  ((∀i.1) 1)

Lemma: fl_all-0
I,J,i:Top.  ((∀i.0) 0)

Lemma: fl_all-fl1
[I:fset(ℕ)]. ∀[i:ℕ]. ∀[x:names(I+i)].  ((∀i.(x=1)) if (x =z i) then else (x=1) fi  ∈ Point(face_lattice(I)))

Lemma: fl_all-fl0
[I:fset(ℕ)]. ∀[i:ℕ]. ∀[x:names(I+i)].  ((∀i.(x=0)) if (x =z i) then else (x=0) fi  ∈ Point(face_lattice(I)))

Lemma: fl_all-id
[I:fset(ℕ)]. ∀[i:{i:ℕ| ¬i ∈ I} ]. ∀[phi:Point(face_lattice(I))].  ((∀i.phi) phi ∈ Point(face_lattice(I)))

Lemma: fl_all-meet
[I:fset(ℕ)]. ∀[i:ℕ]. ∀[a,b:Point(face_lattice(I+i))].  ((∀i.a ∧ b) (∀i.a) ∧ (∀i.b) ∈ Point(face_lattice(I)))

Lemma: fl_all-join
[I:fset(ℕ)]. ∀[i:ℕ]. ∀[a,b:Point(face_lattice(I+i))].  ((∀i.a ∨ b) (∀i.a) ∨ (∀i.b) ∈ Point(face_lattice(I)))

Lemma: fl_all_com
[I:fset(ℕ)]. ∀[i,j:ℕ]. ∀[phi:Point(face_lattice(I+i+j))].  ((∀i.(∀j.phi)) (∀j.(∀i.phi)) ∈ Point(face_lattice(I)))

Lemma: fl_all_decomp
[I:fset(ℕ)]. ∀[i:ℕ]. ∀[phi:Point(face_lattice(I+i))].
  (phi (∀i.phi) ∨ phi ∧ (i=0) ∨ phi ∧ (i=1) ∈ Point(face_lattice(I+i)))

Lemma: fl_all-implies-instance
[I:fset(ℕ)]. ∀[x:Point(dM(I))]. ∀[i:ℕ]. ∀[v:Point(face_lattice(I+i))].
  (((∀i.v) 1 ∈ Point(face_lattice(I)))  ((v)<(i/x)> 1 ∈ Point(face_lattice(I))))

Definition: dM-fl-all
dM-fl-all(I;phi;z) ==
  lattice-extend(face_lattice(I);union-deq(names(I);names(I);NamesDeq;NamesDeq);face_lattice-deq();λx.case x
                                                                                                       of inl(a) =>
                                                                                                       (∀a.phi)
                                                                                                       inr(a) =>
                                                                                                       (∀a.phi);z)

Lemma: dM-fl-all_wf
[I:fset(ℕ)]. ∀[phi:Point(face_lattice(I))]. ∀[z:Point(dM(I))].  (dM-fl-all(I;phi;z) ∈ Point(face_lattice(I)))

Lemma: fl-morph-fl_all
[I,J:fset(ℕ)]. ∀[f:J ⟶ I]. ∀[j:{i:ℕ| ¬i ∈ J} ]. ∀[i:{i:ℕ| ¬i ∈ I} ]. ∀[psi:Point(face_lattice(I+i))].
  (((∀i.psi))<f> (∀j.(psi)<f,i=j>) ∈ Point(face_lattice(J)))

Definition: face-lattice-tube
face-lattice-tube(I;phi;j) ==  s(phi) ∨ (j=0) ∨ (j=1)

Lemma: face-lattice-tube_wf
[I:fset(ℕ)]. ∀[phi:Point(face_lattice(I))]. ∀[j:ℕ].  (face-lattice-tube(I;phi;j) ∈ Point(face_lattice(I+j)))

Lemma: fl-morph-face-lattice-tube1
[J:fset(ℕ)]. ∀[k:names(J)]. ∀[I:fset(ℕ)]. ∀[phi:Point(face_lattice(I))]. ∀[j:ℕ]. ∀[g:J ⟶ I+j].
  (face-lattice-tube(I;phi;j))<g> (s(phi))<g> ∨ (k=0) ∨ (k=1) ∈ Point(face_lattice(J)) 
  supposing (g j) = <k> ∈ Point(dM(J))

Comment: cubical subset doc
================================= cubical subset ===============================
 We next define the subset of the formal cube I,psi  where psi ∈ 𝔽(I).
 Later, we will define more general subset operation G, phi on any context G
 where phi ∈ {G ⊢ _:𝔽is cubical term of type ⌜𝔽⌝⋅

Definition: name-morph-satisfies
(psi f) ==  (psi)<f> 1 ∈ Point(face_lattice(J))

Lemma: name-morph-satisfies_wf
[I,J:fset(ℕ)]. ∀[psi:Point(face_lattice(I))]. ∀[f:J ⟶ I].  ((psi f) 1 ∈ ℙ)

Lemma: name-morph-satisfies-comp
[I,J,K:fset(ℕ)]. ∀[psi:Point(face_lattice(I))]. ∀[f:J ⟶ I]. ∀[g:K ⟶ J].  uiff((f(psi) g) 1;(psi f ⋅ g) 1)

Lemma: implies-nh-comp-satisfies
[I,J,K:fset(ℕ)]. ∀[psi:Point(face_lattice(I))]. ∀[f:J ⟶ I]. ∀[g:K ⟶ J].  (psi f ⋅ g) supposing (psi f) 1

Lemma: name-morph-satisfies-join
I,J:fset(ℕ). ∀a,b:Point(face_lattice(I)). ∀f:J ⟶ I.  ((fl-join(I;a;b) f) ⇐⇒ (a f) 1 ∨ (b f) 1)

Lemma: name-morph-satisfies-fset-join
I,J:fset(ℕ). ∀f:J ⟶ I. ∀s:fset(Point(face_lattice(I))).
  ((\/(s) f) ⇐⇒ ↓∃a:Point(face_lattice(I)). (a ∈ s ∧ (a f) 1))

Lemma: name-morph-satisfies-fl0
[I,J:fset(ℕ)]. ∀[i:names(I)]. ∀[f:J ⟶ I].  uiff(((i=0) f) 1;(f i) 0 ∈ Point(dM(J)))

Lemma: name-morph-satisfies-fl1
[I,J:fset(ℕ)]. ∀[i:names(I)]. ∀[f:J ⟶ I].  uiff(((i=1) f) 1;(f i) 1 ∈ Point(dM(J)))

Lemma: name-morph-satisfies-0
[I:fset(ℕ)]. ∀[i:ℕ].  ((i=0) (i0)) 1

Lemma: name-morph-satisfies-0'
[I:fset(ℕ)]. ∀[i:ℕ].  ((i=0) (i0)) 1

Lemma: name-morph-1-satisfies
[I,J:fset(ℕ)]. ∀[f:J ⟶ I].  (1 f) 1

Lemma: satisfies-irr-face
[I,J:fset(ℕ)]. ∀[as,bs:fset(names(I))]. ∀[g:J ⟶ I].
  uiff((irr_face(I;as;bs) g) 1;(∀a:names(I). (a ∈ as  ((g a) 0 ∈ Point(dM(J)))))
  ∧ (∀b:names(I). (b ∈ bs  ((g b) 1 ∈ Point(dM(J))))))

Lemma: irr-face-morph-satisfies
[I:fset(ℕ)]. ∀[as,bs:fset(names(I))].
  (irr_face(I;as;bs) irr-face-morph(I;as;bs)) supposing ↑fset-disjoint(NamesDeq;as;bs)

Lemma: irr-face-morph-property
[I:fset(ℕ)]. ∀[as,bs:fset(names(I))]. ∀[J:fset(ℕ)]. ∀[g:J ⟶ I].
  ((irr_face(I;as;bs) g)  (g irr-face-morph(I;as;bs) ⋅ g ∈ J ⟶ I))

Lemma: face_lattice-le
[I:fset(ℕ)]. ∀[x,y:Point(face_lattice(I))].
  uiff(x ≤ y;∀f:I ⟶ I. (((x)<f> 1 ∈ Point(face_lattice(I)))  ((y)<f> 1 ∈ Point(face_lattice(I)))))

Lemma: satisfies-face-lattice-tube
I:fset(ℕ). ∀i:{i:ℕ| ¬i ∈ I} . ∀phi:𝔽(I). ∀j:{j:ℕ| ¬j ∈ I+i} . ∀K:fset(ℕ). ∀g:K ⟶ I+j.
  ((face-lattice-tube(I;phi;j) g) ⇐⇒ (phi s ⋅ g) 1 ∨ ((g j) 0 ∈ Point(dM(K))) ∨ ((g j) 1 ∈ Point(dM(K))))

Lemma: satisfies-s-face-lattice-tube
I:fset(ℕ). ∀i:{i:ℕ| ¬i ∈ I} . ∀phi:𝔽(I). ∀j:{j:ℕ| ¬j ∈ I+i} . ∀K:fset(ℕ). ∀f:K ⟶ I+j+i.
  ((s(face-lattice-tube(I;phi;j)) f) 1
  ⇐⇒ (s(phi) s ⋅ f) 1 ∨ ((s ⋅ j) 0 ∈ Point(dM(K))) ∨ ((s ⋅ j) 1 ∈ Point(dM(K))))

Definition: cubical-subset
I,psi ==  rep-sub-sheaf(CubeCat;I;λJ,f. (psi f) 1)

Lemma: cubical-subset_wf
[I:fset(ℕ)]. ∀[psi:𝔽(I)].  I,psi j⊢

Lemma: cubical-subset-I_cube
[I,psi,J:Top].  (I,psi(J) {f:J ⟶ I| (psi f) 1} )

Lemma: member-cubical-subset-I_cube
[I:fset(ℕ)]. ∀[psi:𝔽(I)]. ∀[J:fset(ℕ)]. ∀[f:J ⟶ I].  f ∈ I,psi(J) supposing (psi f) 1

Lemma: cubical-subset-I_cube-member
[I:fset(ℕ)]. ∀[psi:𝔽(I)]. ∀[J:fset(ℕ)]. ∀[f:I,psi(J)].  ((f ∈ J ⟶ I) ∧ (psi f) 1)

Lemma: cubical-subset-restriction
[I,J,K,f,g,phi:Top].  (g(f) f ⋅ g)

Definition: subset-iota
iota ==  λK,g. g

Lemma: subset-iota_wf
[I:fset(ℕ)]. ∀[psi:𝔽(I)].  (iota ∈ I,psi j⟶ formal-cube(I))

Definition: subset_iota
subset_iota(psi) ==  λK,g. g

Lemma: subset_iota_wf
[I:fset(ℕ)]. ∀[psi:𝔽(I)].  (subset_iota(psi) ∈ I,psi ⟶ formal-cube(I))

Definition: subset-trans
subset-trans(I;J;f;x) ==  λK,g. f ⋅ g

Lemma: subset-trans_wf
[I,J:fset(ℕ)]. ∀[f:J ⟶ I]. ∀[psi:𝔽(I)].  (subset-trans(I;J;f;psi) ∈ J,(psi)<f> j⟶ I,psi)

Lemma: subset-trans-iota-lemma
Gamma:j⊢. ∀I:fset(ℕ). ∀rho:Gamma(I). ∀phi:𝔽(I). ∀J:fset(ℕ). ∀f:J ⟶ I.
  (<rho> iota subset-trans(I;J;f;phi) = <f(rho)> iota ∈ J,(phi)<f> j⟶ Gamma)

Comment: Rules of MLTT without type formers
The rules from Figure in Bezem, Coquand, Huber
"A model of type theory in cubical sets".
But now we interpret everything in the new cubical type theory.
(First those without type formers: 20 lemmas)

    Γ ⊢
-------------                         this is csm-id_wf
  1: Γ ⟶ Γ

σ: Δ ⟶ Γ    δ: Θ -> Δ
------------------------              this is csm-comp_wf
     σ δ : Θ ⟶ Γ

Γ ⊢ A   σ : Δ ⟶ Γ
----------------------                this is csm-ap-type_wf
     Δ ⊢ A σ

Γ ⊢ t: A   σ: Δ ⟶ Γ
-----------------------               this is csm-ap-term_wf
     Δ ⊢ t σA σ


-------------                         this is trivial-cube-set_wf
   () ⊢

Γ ⊢    Γ ⊢ A
-----------------                     this is cube-context-adjoin_wf
     Γ.A ⊢ 

     Γ ⊢ A
------------------                    this is cc-fst_wf
   p:  Γ.A ⟶ Γ

   Γ ⊢ A
-----------------                     this is cc-snd_wf
   Γ.A ⊢ q: Ap 

σ:Δ ⟶ Γ  Γ ⊢ A   Δ ⊢ u:A s
------------------------------        this is csm-adjoin_wf
   ,u) ⊢ Δ ⟶ Γ.A

1 σ = σ = σ                         this is csm-id-comp
(σ δ) ν = σ (δ ν)                     this is csm-comp-assoc
[u] (1,u)                           this is csm-id-adjoin
A                               this is csm-ap-id-type
(A σ) δ (σ δ)                     this is csm-ap-comp-type
u                               this is csm-ap-id-term
(u σ) δ (σ δ)                     this is csm-ap-comp-term 
u) δ (σ δu δ)                 this is csm-adjoin-ap
u) = σ                          this is cc-fst-csm-adjoin
u) u                          this is cc-snd-csm-adjoin
(p,q) 1                             this is csm-adjoin-fst-snd



Comment: Rules of MLTT type formers
The rules about type formers Π and ⌜Σ⌝ from Figure in Bezem, Coquand, Huber
"A model of type theory in cubical sets".
(Interpreted in the new cubical type theory: 19 lemmas)

  Γ.A ⊢ B
-------------                         this is cubical-pi_wf
 Γ ⊢ Π 

  Γ.A ⊢ B   Γ.A ⊢ b:B
------------------------              this is cubical-lambda_wf
   Γ ⊢ λ: Π B  

  Γ.A ⊢ B
-------------                         this is cubical-sigma_wf
 Γ ⊢ Σ B


 Γ.A ⊢ B  Γ ⊢ u:A  Γ ⊢ v:B[u]
------------------------------        this is cubical-pair_wf
     Γ ⊢ (u,v): Σ B

  Γ ⊢ w: Σ B
----------------                      this is cubical-fst_wf
  Γ ⊢ w.1: A

  Γ ⊢ w: Σ B
-----------------                     this is cubical-snd_wf
  Γ ⊢ w.2: B[w.1]   

 Γ ⊢ w:Π B   Γ ⊢ u:A
------------------------              this is cubical-app_wf
 Γ ⊢ app(w,u): B[u]  


(Π B)σ = Π (Aσ(B(σp,q))           this is csm-cubical-pi
b)σ = λ(b(σp,q))                    this is csm-ap-cubical-lambda  
app(w,u)δ app(wδ)               this is csm-ap-cubical-app
app(λb, u) b[u]                     this is cubical-beta
= λ(app(wp,q))                      this is cubical-eta
(Σ B)σ =  (Aσ(B(σp,q))            this is csm-cubical-sigma
(w.1)σ (wσ).1                       this is csm-ap-cubical-fst
(w.2)σ (wσ).2                       this is csm-ap-cubical-snd
(u,v)σ (uσ)                     this is csm-ap-cubical-pair
(u,v).1 u                           this is cubical-fst-pair
(u,v).2 v                           this is cubical-snd-pair
(w.1,w.2) w                         this is cubical-pair-eta


Definition: cubical_type
cubical_type{i:l}(X) ==  presheaf{i':l}(cubes(X))

Lemma: cubical_type_wf
[X:j⊢]. (cubical_type{i:l}(X) ∈ 𝕌{[i'' j']})

Definition: cubical-type
{X ⊢ _} ==
  {AF:A:I:fset(ℕ) ⟶ X(I) ⟶ Type × (I:fset(ℕ) ⟶ J:fset(ℕ) ⟶ f:J ⟶ I ⟶ a:X(I) ⟶ (A a) ⟶ (A f(a)))| 
   let A,F AF 
   in (∀I:fset(ℕ). ∀a:X(I). ∀u:A a.  ((F u) u ∈ (A a)))
      ∧ (∀I,J,K:fset(ℕ). ∀f:J ⟶ I. ∀g:K ⟶ J. ∀a:X(I). ∀u:A a.
           ((F f ⋅ u) (F f(a) (F u)) ∈ (A f ⋅ g(a))))} 

Lemma: cubical-type_wf
[X:j⊢]. (X ⊢  ∈ 𝕌{[i' j']})

Lemma: cubical-type-sq-presheaf-type
[X:Top]. ({X ⊢ _} {X ⊢ _})

Lemma: cubical-type-equal
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[B:A:I:fset(ℕ) ⟶ X(I) ⟶ Type × (I:fset(ℕ)
                                                         ⟶ J:fset(ℕ)
                                                         ⟶ f:J ⟶ I
                                                         ⟶ a:X(I)
                                                         ⟶ (A a)
                                                         ⟶ (A f(a)))].
  B ∈ {X ⊢ _} 
  supposing A
  B
  ∈ (A:I:fset(ℕ) ⟶ X(I) ⟶ Type × (I:fset(ℕ) ⟶ J:fset(ℕ) ⟶ f:J ⟶ I ⟶ a:X(I) ⟶ (A a) ⟶ (A f(a))))

Lemma: cubical-type-equal2
[X:j⊢]. ∀[A,B:{X ⊢ _}].
  B ∈ {X ⊢ _} 
  supposing A
  B
  ∈ (A:I:fset(ℕ) ⟶ X(I) ⟶ Type × (I:fset(ℕ) ⟶ J:fset(ℕ) ⟶ f:J ⟶ I ⟶ a:X(I) ⟶ (A a) ⟶ (A f(a))))

Definition: cubical-type-iso
cubical-type-iso(X) ==  λF.<λI,rho. (F <I, rho>), λI,J,f,a. (F <I, a> <J, f(a)> f)>

Lemma: cubical-type-iso_wf
[X:j⊢]. (cubical-type-iso(X) ∈ cubical_type{i:l}(X) ⟶ {X ⊢_})

Definition: cubical-type-rev-iso
cubical-type-rev-iso(X) ==
  λF.Presheaf(Set(p) =(fst(F)) (fst(p)) (snd(p));
              Morphism(p,q,f,a) (snd(F)) (fst(p)) (fst(q)) (snd(p)) a)

Lemma: cubical-type-rev-iso_wf
[X:j⊢]. (cubical-type-rev-iso(X) ∈ {X ⊢ _} ⟶ cubical_type{i:l}(X))

Lemma: cubical-type-iso-inverse
[X:j⊢]. ((cubical-type-rev-iso(X) cubical-type-iso(X)) x.x) ∈ (cubical_type{i:l}(X) ⟶ cubical_type{i:l}(X)))

Lemma: cubical-type-iso-inverse2
[X:j⊢]. ((cubical-type-iso(X) cubical-type-rev-iso(X)) x.x) ∈ ({X ⊢ _} ⟶ {X ⊢ _}))

Lemma: cubical-type-cumulativity2
[X:j⊢]. ({X ⊢ _} ⊆cubical-type{[j i]:l}(X))

Lemma: cubical-type-cumulativity
[X:j⊢]. ({X ⊢ _} ⊆{X ⊢_})

Lemma: cubical-type-cumulativity-i-j
[X:j⊢]. ({X ⊢_} ⊆cubical-type{[j i]:l}(X))

Definition: cubical-type-at
A(a) ==  (fst(A)) a

Lemma: cubical-type-at_wf1
[X:j⊢]. ∀[A:{X ⊢_}]. ∀[I:fset(ℕ)]. ∀[a:X(I)].  (A(a) ∈ 𝕌{j})

Lemma: cubical-type-at_wf
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[I:fset(ℕ)]. ∀[a:X(I)].  (A(a) ∈ Type)

Lemma: istype-cubical-type-at
[X:j⊢]. ∀[I:fset(ℕ)]. ∀[a:X(I)]. ∀[A:{X ⊢ _}].  istype(A(a))

Lemma: cubical_type_at_pair_lemma
a,I,B,A:Top.  (<A, B>(a) a)

Definition: cubical-type-ap-morph
(u f) ==  (snd(A)) u

Lemma: cubical-type-ap-morph_wf
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[I,J:fset(ℕ)]. ∀[f:J ⟶ I]. ∀[a:X(I)]. ∀[u:A(a)].  ((u f) ∈ A(f(a)))

Lemma: cubical_type_ap_morph_pair_lemma
u,a,f,J,I,G,F:Top.  ((u f) u)

Lemma: cubical-type-ap-morph-comp-general
[X:j⊢]. ∀[A:{X ⊢_}]. ∀[I,J,K:fset(ℕ)]. ∀[f:J ⟶ I]. ∀[g:K ⟶ J]. ∀[a:X(I)]. ∀[u:A(a)].
  (((u f) f(a) g) (u f ⋅ g) ∈ A(f ⋅ g(a)))

Lemma: cubical-type-ap-morph-comp
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[I,J,K:fset(ℕ)]. ∀[f:J ⟶ I]. ∀[g:K ⟶ J]. ∀[a:X(I)]. ∀[u:A(a)].
  (((u f) f(a) g) (u f ⋅ g) ∈ A(f ⋅ g(a)))

Lemma: cubical-type-ap-morph-comp-eq-general
[X:j⊢]. ∀[A:{X ⊢_}]. ∀[I,J,K:fset(ℕ)]. ∀[f:J ⟶ I]. ∀[g:K ⟶ J]. ∀[a:X(I)]. ∀[b:X(J)]. ∀[u:A(a)].
  ((u f) g) (u f ⋅ g) ∈ A(f ⋅ g(a)) supposing f(a) ∈ X(J)

Lemma: cubical-type-ap-morph-comp-eq
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[I,J,K:fset(ℕ)]. ∀[f:J ⟶ I]. ∀[g:K ⟶ J]. ∀[a:X(I)]. ∀[b:X(J)]. ∀[u:A(a)].
  ((u f) g) (u f ⋅ g) ∈ A(f ⋅ g(a)) supposing f(a) ∈ X(J)

Lemma: cubical-type-ap-morph-id
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[I:fset(ℕ)]. ∀[f:I ⟶ I]. ∀[a:X(I)]. ∀[u:A(a)].  (u f) u ∈ A(a) supposing 1 ∈ I ⟶ I

Lemma: cubical-type-equal3
[X:j⊢]. ∀[A,B:{X ⊢ _}].
  (A B ∈ {X ⊢ _}) supposing 
     ((∀I:fset(ℕ). ∀[rho:X(I)]. ∀[J:fset(ℕ)]. ∀[f:J ⟶ I]. ∀[u:A(rho)].  ((u rho f) (u rho f) ∈ A(f(rho)))) and 
     (∀I:fset(ℕ). ∀[rho:X(I)]. (A(rho) B(rho) ∈ Type)))

Definition: csm-ap-type
(AF)s ==  let A,F AF in <λI,a. (A (s)a), λI,J,f,a,u. (F (s)a u)>

Lemma: csm-ap-type_wf1
[Gamma,Delta:j⊢]. ∀[A:{Gamma ⊢_}]. ∀[s:Delta j⟶ Gamma].  Delta ⊢(A)s

Lemma: csm-ap-type_wf
[Gamma,Delta:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[s:Delta j⟶ Gamma].  Delta ⊢ (A)s

Lemma: csm-ap-type-subset-iota
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[s:Top].  (((A)s)iota (A)s)

Definition: csm-type-ap
csm-type-ap(A;s) ==  (A)s

Lemma: csm-type-ap_wf
[Gamma,Delta:j⊢]. ∀[s:Delta j⟶ Gamma]. ∀[A:{Gamma ⊢ _}].  Delta ⊢ csm-type-ap(A;s)

Lemma: context-map-ap-type
[I:fset(ℕ)]. ∀[Gamma:j⊢]. ∀[rho:Gamma(I)]. ∀[A:{Gamma ⊢ _}].  formal-cube(I) ⊢ (A)<rho>

Lemma: csm-ap-type-at
[B,s,x,K:Top].  ((B)s(x) B((s)x))

Lemma: csm-ap-id-type
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}].  ((A)1(Gamma) A ∈ {Gamma ⊢ _})

Lemma: csm-ap-type-is-id
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[s:Gamma j⟶ Gamma].  (A)s A ∈ {Gamma ⊢ _} supposing 1(Gamma) ∈ Gamma j⟶ Gamma

Lemma: csm-ap-type-iota
[X:j⊢]. ∀[A:{X ⊢ _}].  ((A)iota A ∈ {X ⊢ _})

Lemma: csm-comp-type
[Gamma,Delta,Z,s1,s2,A:Top].  ((A)s2 s1 ((A)s2)s1)

Lemma: csm-type-comp
[Gamma,A,Delta,Z,s1,s2:Top].  (((A)s2)s1 (A)s2 s1)

Lemma: csm-ap-comp-type
[Gamma,Delta,Z:j⊢]. ∀[s1:Z j⟶ Delta]. ∀[s2:Delta j⟶ Gamma]. ∀[A:{Gamma ⊢ _}].  ((A)s2 s1 ((A)s2)s1 ∈ {Z ⊢ _})

Lemma: csm-ap-comp-type-sq
[Gamma,Delta,Z,s1,s2,A:Top].  ((A)s2 s1 ((A)s2)s1)

Lemma: csm-ap-comp-type-sq2
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[s1,s2:Top].  (((A)s2)s1 (A)s2 s1)

Lemma: csm-cubical-type-ap-morph
[A,I,J,f,a,u,s:Top].  ((u f) (u (s)a f))

Definition: closed-cubical-type
* ⊢ _} ==
  {AF:A:I:fset(ℕ) ⟶ Type × (I:fset(ℕ) ⟶ J:fset(ℕ) ⟶ f:J ⟶ I ⟶ (A I) ⟶ (A J))| 
   let A,F AF 
   in (∀I:fset(ℕ). ∀u:A I.  ((F u) u ∈ (A I)))
      ∧ (∀I,J,K:fset(ℕ). ∀f:J ⟶ I. ∀g:K ⟶ J. ∀u:A I.  ((F f ⋅ u) (F (F u)) ∈ (A K)))} 

Lemma: closed-cubical-type_wf
* ⊢ _} ∈ 𝕌'

Lemma: closed-cubical-type-cumulativity
* ⊢ _} ⊆* ⊢_}

Definition: closed-type-to-type
closed-type-to-type(T) ==  let A,F in <λI,a. (A I), λI,J,f,a,u. (F u)>

Lemma: closed-type-to-type_wf
[T:{ * ⊢ _}]. ∀[X:j⊢].  X ⊢ closed-type-to-type(T)

Lemma: csm-closed-type
[T:{ * ⊢ _}]. ∀[s:Top].  ((closed-type-to-type(T))s closed-type-to-type(T))

Definition: cubical-term
{X ⊢ _:A} ==  {u:I:fset(ℕ) ⟶ a:X(I) ⟶ A(a)| ∀I,J:fset(ℕ). ∀f:J ⟶ I. ∀a:X(I).  ((u f) (u f(a)) ∈ A(f(a)))} 

Lemma: cubical-term_wf
[X:j⊢]. ∀[A:{X ⊢ _}].  ({X ⊢ _:A} ∈ 𝕌{[i j']})

Lemma: cubical-term-eqcd
[X:j⊢]. ∀[A,B:{X ⊢ _}].  {X ⊢ _:A} {X ⊢ _:B} ∈ 𝕌{[i j']} supposing B ∈ {X ⊢ _}

Lemma: istype-cubical-term
[X:j⊢]. ∀[A:{X ⊢ _}].  istype({X ⊢ _:A})

Lemma: istype-cubical-term-closed-type
[X:j⊢]. ∀[T:{ * ⊢ _}].  istype({X ⊢ _:closed-type-to-type(T)})

Lemma: cubical-term-sq-presheaf-term
[X,A:Top].  ({X ⊢ _:A} {X ⊢ _:A})

Definition: cubical-term-at
u(a) ==  a

Lemma: cubical-term-at_wf1
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[u:{X ⊢ _:A}]. ∀[I:fset(ℕ)]. ∀[a:X(I)].  (u(a) ∈ A(a))

Lemma: cubical-term-at_wf
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[u:{X ⊢ _:A}]. ∀[I:fset(ℕ)]. ∀[a:X(I)].  (u(a) ∈ A(a))

Lemma: cubical-term-at-morph1
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[u:{X ⊢ _:A}]. ∀[I:fset(ℕ)]. ∀[a:X(I)]. ∀[J:fset(ℕ)]. ∀[f:J ⟶ I].
  ((u(a) f) u(f(a)) ∈ A(f(a)))

Lemma: cubical-term-at-morph
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[u:{X ⊢ _:A}]. ∀[I:fset(ℕ)]. ∀[a:X(I)]. ∀[J:fset(ℕ)]. ∀[f:J ⟶ I].
  ((u(a) f) u(f(a)) ∈ A(f(a)))

Lemma: cubical-term-equal
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[u:{X ⊢ _:A}]. ∀[z:I:fset(ℕ) ⟶ a:X(I) ⟶ A(a)].
  z ∈ {X ⊢ _:A} supposing z ∈ (I:fset(ℕ) ⟶ a:X(I) ⟶ A(a))

Lemma: cubical-term-equal2
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[u,z:{X ⊢ _:A}].  z ∈ {X ⊢ _:A} supposing ∀I:fset(ℕ). ∀a:X(I).  ((u a) (z a) ∈ A(a))

Definition: closed-cubical-term
closed-cubical-term(T) ==
  {u:I:fset(ℕ) ⟶ ((fst(T)) I)| ∀I,J:fset(ℕ). ∀f:J ⟶ I.  (((snd(T)) (u I)) (u J) ∈ ((fst(T)) J))} 

Lemma: closed-cubical-term_wf
[T:{ * ⊢ _}]. (closed-cubical-term(T) ∈ Type)

Definition: closed-term-to-term
closed-term-to-term(t) ==  λI,a. (t I)

Lemma: closed-term-to-term_wf
[T:{ * ⊢ _}]. ∀[t:closed-cubical-term(T)]. ∀[X:j⊢].  (closed-term-to-term(t) ∈ {X ⊢ _:closed-type-to-type(T)})

Definition: canonical-section
canonical-section(Gamma;A;I;rho;a) ==  λK,f. (a rho f)

Lemma: canonical-section_wf
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[I:fset(ℕ)]. ∀[rho:Gamma(I)]. ∀[a:A(rho)].
  (canonical-section(Gamma;A;I;rho;a) ∈ {formal-cube(I) ⊢ _:(A)<rho>})

Lemma: canonical-section-at
[Gamma,A,I,rho,a,J,x:Top].  (canonical-section(Gamma;A;I;rho;a)(x) (a rho x))

Definition: trivial-section
The trivial section, () must satisfy I,0 |- ():A subset-iota().
But since the "cubical subset" I,0 is empty, everything satisfies this
property; so we could define it to be anything we like.
(We chose ⌜⋅⌝.)⋅

() ==  ⋅

Lemma: trivial-section_wf
[I:fset(ℕ)]. ∀[X:Top]. ∀[phi:Point(face_lattice(I))].  () ∈ {I,phi ⊢ _:X} supposing phi 0 ∈ Point(face_lattice(I))

Definition: csm-ap-term
(t)s ==  λI,a. (t (s)a)

Lemma: csm-ap-term_wf1
[Delta,Gamma:j⊢]. ∀[A:{Gamma ⊢_}]. ∀[s:Delta j⟶ Gamma]. ∀[t:{Gamma ⊢ _:A}].  ((t)s ∈ {Delta ⊢ _:(A)s})

Lemma: csm-ap-term_wf
[Delta,Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[s:Delta j⟶ Gamma]. ∀[t:{Gamma ⊢ _:A}].  ((t)s ∈ {Delta ⊢ _:(A)s})

Lemma: csm-ap-term-at
[J,s,rho,t:Top].  ((t)s(rho) t((s)rho))

Lemma: csm-ap-id-term
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[t:{Gamma ⊢ _:A}].  ((t)1(Gamma) t ∈ {Gamma ⊢ _:A})

Lemma: csm-comp-term
[Gamma,Delta,Z,s1,s2,t:Top].  ((t)s2 s1 ((t)s2)s1)

Lemma: csm-ap-comp-term-sq
[Gamma,Delta,Z,s1,s2,t:Top].  ((t)s2 s1 ((t)s2)s1)

Lemma: csm-ap-comp-term-sq2
[s1,s2,t:Top].  (((t)s2)s1 (t)s2 s1)

Lemma: csm-ap-comp-term
[Gamma,Delta,Z:j⊢]. ∀[s1:Z j⟶ Delta]. ∀[s2:Delta j⟶ Gamma]. ∀[A:{Gamma ⊢ _}]. ∀[t:{Gamma ⊢ _:A}].
  ((t)s2 s1 ((t)s2)s1 ∈ {Z ⊢ _:(A)s2 s1})

Lemma: context-map-comp2
[G:j⊢]. ∀[I,J:fset(ℕ)]. ∀[f:J ⟶ I]. ∀[a:G(I)].  (<a> o <f> = <f(a)> ∈ formal-cube(J) ij⟶ G)

Lemma: csm-comp-context-map
[Gamma,Delta:j⊢]. ∀[sigma:Delta j⟶ Gamma]. ∀[I:fset(ℕ)]. ∀[rho:Delta(I)].
  (sigma o <rho> = <(sigma)rho> ∈ formal-cube(I) j⟶ Gamma)

Definition: cube-context-adjoin
X.A ==  <λI.(alpha:X(I) × A(alpha)), λI,J,f,p. <f(fst(p)), (snd(p) fst(p) f)>>

Lemma: cube-context-adjoin-wf
[Gamma:j⊢]. ∀[A:{Gamma ⊢_}].  Gamma.A ij⊢

Lemma: cube-context-adjoin_wf
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}].  Gamma.A ij⊢

Lemma: cube-context-adjoin-I_cube
[Gamma,A,I:Top].  (Gamma.A(I) alpha:Gamma(I) × A(alpha))

Definition: cc-adjoin-cube
(v;u) ==  <v, u>

Lemma: cc-adjoin-cube_wf
X:j⊢. ∀A:{X ⊢ _}. ∀J:fset(ℕ). ∀v:X(J). ∀u:A(v).  ((v;u) ∈ X.A(J))

Lemma: cc-adjoin-cube-restriction
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[J,K,g,v,u:Top].  (g((v;u)) (g(v);(u g)))

Definition: cc-fst
==  λI,p. (fst(p))

Lemma: cc-fst_wf
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}].  (p ∈ Gamma.A ij⟶ Gamma)

Definition: cc-snd
==  λI,p. (snd(p))

Lemma: cc-snd_wf
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}].  (q ∈ {Gamma.A ⊢ _:(A)p})

Lemma: cc_fst_adjoin_cube_lemma
y,x,J:Top.  ((p)(x;y) x)

Lemma: cc_snd_adjoin_cube_lemma
y,x,J:Top.  ((q)(x;y) y)

Lemma: csm_comp_fst_adjoin_cube_lemma
y,x,K,s,X,B,A:Top.  ((s p)(x;y) (s)x)

Lemma: csm_comp_snd_adjoin-cube_lemma
y,x,K,s,X,B,A:Top.  ((s q)(x;y) (s)y)

Definition: cc-fstfst
pp ==  λI,rho. (fst(fst(rho)))

Lemma: cc-fstfst_wf
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[B:{Gamma.A ⊢ _}].  (pp ∈ Gamma.A.B ij⟶ Gamma)

Definition: typed-cc-fst
tp{i:l} ==  p

Lemma: typed-cc-fst_wf
[G:j⊢]. ∀[A:{G ⊢ _}].  (tp{i:l} ∈ G.A ij⟶ G)

Definition: typed-cc-snd
tq ==  q

Lemma: typed-cc-snd_wf
G:j⊢. ∀A:{G ⊢ _}.  (tq ∈ {G.A ⊢ _:(A)tp{i:l}})

Definition: cc-m2
q2 ==  (q)p

Lemma: cc-m2_wf
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[B:{X.A ⊢ _}].  (q2 ∈ {X.A.B ⊢ _:((A)p)p})

Definition: cc-m3
q3 ==  (q2)p

Lemma: cc-m3_wf
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[B:{X.A ⊢ _}]. ∀[C:{X.A.B ⊢ _}].  (q3 ∈ {X.A.B.C ⊢ _:(((A)p)p)p})

Definition: cc-m4
q4 ==  (q3)p

Lemma: cc-m4_wf
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[B:{X.A ⊢ _}]. ∀[C:{X.A.B ⊢ _}]. ∀[D:{X.A.B.C ⊢ _}].  (q4 ∈ {X.A.B.C.D ⊢ _:((((A)p)p)p)p})

Lemma: csm-ap-term-p
[Gamma:j⊢]. ∀[A,T:{Gamma ⊢ _}]. ∀[t:{Gamma ⊢ _:A}].  ((t)p ∈ {Gamma.T ⊢ _:(A)p})

Lemma: csm-type-at
s,A,I,alpha:Top.  ((A)s(alpha) A((s)alpha))

Definition: csm-adjoin
(s;u) ==  λI,a. <(s)a, a>

Lemma: csm-adjoin-wf
[Gamma,Delta:j⊢]. ∀[A:{Gamma ⊢_}]. ∀[sigma:Delta j⟶ Gamma]. ∀[u:{Delta ⊢ _:(A)sigma}].
  ((sigma;u) ∈ Delta ij⟶ Gamma.A)

Lemma: csm-adjoin_wf
[Gamma,Delta:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[sigma:Delta j⟶ Gamma]. ∀[u:{Delta ⊢ _:(A)sigma}].
  ((sigma;u) ∈ Delta ij⟶ Gamma.A)

Lemma: csm-adjoin-comp
[Gamma,Delta,X,A,sigma,u,g:Top].  ((sigma;u) (sigma g;(u)g))

Lemma: csm-adjoin-ap
[sigma,u,I,del:Top].  (((sigma;u))del ((sigma)del;(u)del))

Definition: csm-id-adjoin
[u] ==  (1(X);u)

Lemma: csm-id-adjoin-wf
[Gamma:j⊢]. ∀[A:{Gamma ⊢_}]. ∀[u:{Gamma ⊢ _:A}].  ([u] ∈ Gamma ij⟶ Gamma.A)

Lemma: csm-id-adjoin_wf
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[u:{Gamma ⊢ _:A}].  ([u] ∈ Gamma ij⟶ Gamma.A)

Lemma: csm-id-adjoin-ap
X,u,I,a:Top.  (([u])a (a;u a))

Lemma: cc-fst-csm-adjoin-sq
[Gamma,Delta,X,sigma,u:Top].  (p (sigma;u) sigma 1(Delta))

Lemma: cc-fst-csm-adjoin
[Gamma,Delta:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[sigma:Delta j⟶ Gamma]. ∀[u:{Delta ⊢ _:(A)sigma}].
  (p (sigma;u) sigma ∈ Delta j⟶ Gamma)

Lemma: cc_snd_csm_id_adjoin_lemma
u,G:Top.  ((q)[u] (u)1(G))

Lemma: typed_cc_snd_id_adjoin_lemma
A,G,u,X:Top.  ((tq)[u] (u)1(X))

Lemma: csm_id_adjoin_fst_type_lemma
A,t,X:Top.  (((A)p)[t] (A)1(X))

Lemma: csm_id_adjoin_fst_term_lemma
a,t,X:Top.  (((a)p)[t] (a)1(X))

Lemma: csm_id_ap_term_lemma
t,X,s:Top.  (((t)1(X))s (t)s)

Lemma: cc-snd-csm-adjoin-sq
[G,sigma,u:Top].  ((q)(sigma;u) (u)1(G))

Lemma: cc-snd-csm-adjoin
[Gamma,Delta:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[sigma:Delta j⟶ Gamma]. ∀[u:{Delta ⊢ _:(A)sigma}].
  ((q)(sigma;u) u ∈ {Delta ⊢ _:(A)sigma})

Lemma: csm-adjoin-fst-snd
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}].  ((p;q) 1(Gamma.A) ∈ Gamma.A ij⟶ Gamma.A)

Lemma: csm-ap-type-fst-adjoin
[X:j⊢]. ∀[B:{X ⊢ _}]. ∀[s,u:Top].  (((B)p)(s;u) (B)s)

Lemma: csm-ap-type-fst-id-adjoin
[X:j⊢]. ∀[B:{X ⊢ _}]. ∀[u:Top].  (((B)p)[u] B ∈ {X ⊢ _})

Lemma: csm-id-adjoin-ap-type
Gamma,Delta:j⊢. ∀A:{Gamma ⊢ _}. ∀B:{Gamma.A ⊢ _}. ∀sigma:Delta j⟶ Gamma. ∀u:{Delta ⊢ _:(A)sigma}.
  (((B)(sigma p;q))[u] (B)(sigma;u) ∈ {Delta ⊢ _})

Lemma: csm_ap_term_fst_adjoin_lemma
zz,yy,xx:Top.  (((zz)p)(xx;yy) (zz)xx)

Lemma: csm-ap-term-snd-adjoin
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[u:{X ⊢ _:A}]. ∀[xx:Top].  ((q)(xx;u) u ∈ {X ⊢ _:A})

Lemma: csm-adjoin-id-adjoin
[B,xx,s,X,Y,Z:Top].  (((B)(s p;q))[xx] (B)(s;xx))

Definition: csm+
tau+ ==  (tau p;q)

Lemma: csm+_wf
[H,K:j⊢]. ∀[A:{H ⊢ _}]. ∀[tau:K j⟶ H].  (tau+ ∈ K.(A)tau ij⟶ H.A)

Lemma: csm+-comp-csm+-sq
[H,K,X,A,tau,s:Top].  (tau+ s+ tau s+)

Lemma: csm+-id
[G:j⊢]. ∀[A:{G ⊢ _}].  (1(G)+ 1(G.A) ∈ G.A ij⟶ G.A)

Lemma: csm+-comp-csm+
[H,K,X:j⊢]. ∀[A:{H ⊢ _}]. ∀[tau:K j⟶ H]. ∀[s:X j⟶ K].  (tau+ s+ tau s+ ∈ X.((A)tau)s ij⟶ H.A)

Lemma: csm+_wf+
[H,K:j⊢]. ∀[A:{H ⊢ _}]. ∀[tau:K ij⟶ H]. ∀[B:{H.A ⊢ _}].  (tau++ ∈ K.(A)tau.(B)tau+ ij⟶ H.A.B)

Lemma: p-csm+-type
[H,K,A,B,tau:Top].  (((A)p)tau+ ((A)tau)p)

Lemma: csm+-p-type
[C,H,K,A,B,tau:Top].  (((A)p)(tau+ p;q) ((A)tau+)p)

Lemma: p-csm+-term
[H,K,A,t,tau:Top].  (((t)p)tau+ ((t)tau)p)

Lemma: csm+-p-term
[C,H,K,t,B,tau:Top].  (((t)p)(tau+ p;q) ((t)tau+)p)

Lemma: q-csm+
[H,K,A,tau:Top].  ((q)tau+ q)

Lemma: csm-adjoin-p-q
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[B:{X.A ⊢ _}].  ((B)(p;q) B ∈ {X.A ⊢ _})

Definition: csm-swap
csm-swap(G;A;B) ==  (p+;(q)p)

Lemma: csm-swap_wf
[G:j⊢]. ∀[A,B:{G ⊢ _}].  (csm-swap(G;A;B) ∈ G.A.(B)p ij⟶ G.B.(A)p)

Definition: cubical-pi-family
cubical-pi-family(X;A;B;I;a) ==
  {w:J:fset(ℕ) ⟶ f:J ⟶ I ⟶ u:A(f(a)) ⟶ B((f(a);u))| 
   ∀J,K:fset(ℕ). ∀f:J ⟶ I. ∀g:K ⟶ J. ∀u:A(f(a)).  ((w (f(a);u) g) (w f ⋅ (u f(a) g)) ∈ B(g((f(a);u))))} 

Lemma: cubical-pi-family_wf
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[B:{X.A ⊢ _}]. ∀[I:fset(ℕ)]. ∀[a:X(I)].  (cubical-pi-family(X;A;B;I;a) ∈ Type)

Lemma: cubical-pi-family-comp
X,Delta:j⊢. ∀s:Delta j⟶ X. ∀I,J:fset(ℕ). ∀f:J ⟶ I. ∀a:Delta(I). ∀A:{X ⊢ _}. ∀B:{X.A ⊢ _}.
w:cubical-pi-family(X;A;B;I;(s)a).
  K,g. (w f ⋅ g) ∈ cubical-pi-family(X;A;B;J;(s)f(a)))

Lemma: csm-cubical-pi-family
X,Delta:j⊢. ∀A:{X ⊢ _}. ∀B:{X.A ⊢ _}. ∀s:Delta j⟶ X. ∀I:fset(ℕ). ∀a:Delta(I).
  (cubical-pi-family(X;A;B;I;(s)a) cubical-pi-family(Delta;(A)s;(B)(s p;q);I;a) ∈ Type)

Definition: cubical-fun-family
cubical-fun-family(X; A; B; I; a) ==
  {w:J:fset(ℕ) ⟶ f:J ⟶ I ⟶ u:A(f(a)) ⟶ B(f(a))| 
   ∀J,K:fset(ℕ). ∀f:J ⟶ I. ∀g:K ⟶ J. ∀u:A(f(a)).  ((w f(a) g) (w f ⋅ (u f(a) g)) ∈ B(g(f(a))))} 

Lemma: cubical-fun-family_wf
[X:j⊢]. ∀[A,B:{X ⊢ _}]. ∀[I:fset(ℕ)]. ∀[a:X(I)].  (cubical-fun-family(X; A; B; I; a) ∈ Type)

Lemma: cubical-fun-family-comp
X,Delta:j⊢. ∀s:Delta j⟶ X. ∀I,J:fset(ℕ). ∀f:J ⟶ I. ∀a:Delta(I). ∀A,B:{X ⊢ _}.
w:cubical-fun-family(X; A; B; I; (s)a).
  K,g. (w f ⋅ g) ∈ cubical-fun-family(X; A; B; J; (s)f(a)))

Lemma: csm-cubical-fun-family
X,Delta:j⊢. ∀A,B:{X ⊢ _}. ∀s:Delta j⟶ X. ∀I:fset(ℕ). ∀a:Delta(I).
  (cubical-fun-family(X; A; B; I; (s)a) cubical-fun-family(Delta; (A)s; (B)s; I; a) ∈ Type)

Definition: cubical-pi
Π==  <λI,a. cubical-pi-family(X;A;B;I;a), λI,J,f,a,w,K,g. (w f ⋅ g)>

Lemma: cubical-pi_wf
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[B:{X.A ⊢ _}].  X ⊢ ΠB

Lemma: csm-cubical-pi
X,Delta:j⊢. ∀A:{X ⊢ _}. ∀B:{X.A ⊢ _}. ∀s:Delta j⟶ X.  ((ΠB)s Delta ⊢ Π(A)s (B)(s p;q) ∈ {Delta ⊢ _})

Lemma: cubical-pi-p
X:j⊢. ∀T,A:{X ⊢ _}. ∀B:{X.A ⊢ _}.  ((ΠB)p X.T ⊢ Π(A)p (B)(p p;q) ∈ {X.T ⊢ _})

Definition: csm-dependent
(s)dep ==  (s tp{i:l};tq)

Lemma: csm-dependent_wf
X,Delta:j⊢. ∀A:{X ⊢ _}. ∀s:Delta j⟶ X.  ((s)dep ∈ Delta.(A)s j⟶ X.A)

Lemma: csm_dep_p_rewrite_lemma
T,s,A,Delta,X:Top.  (((T)p)(s)dep ((T)s)p)

Lemma: csm_dep_typed_p_lemma
T,s,A,Delta,X:Top.  (((T)tp{i:l})(s)dep ((T)s)tp{i:l})

Lemma: csm-dep_term_p_lemma
t,s,A,Delta,X:Top.  (((t)tp{i:l})(s)dep ((t)s)tp{i:l})

Lemma: csm_dep_term_q_lemma
s,A,Delta,X:Top.  ((tq)(s)dep tq)

Lemma: csm-cubical-pi-typed
X,Delta:j⊢. ∀A:{X ⊢ _}. ∀B:{X.A ⊢ _}. ∀s:Delta ij⟶ X.  ((ΠB)s Delta ⊢ Π(A)s (B)(s)dep ∈ {Delta ⊢ _})

Definition: cubical-fun
(A ⟶ B) ==  <λI,a. cubical-fun-family(X; A; B; I; a), λI,J,f,a,w,K,g. (w f ⋅ g)>

Lemma: cubical-fun_wf
[X:j⊢]. ∀[A,B:{X ⊢ _}].  X ⊢ (A ⟶ B)

Lemma: csm-cubical-fun
X,Delta:j⊢. ∀A,B:{X ⊢ _}. ∀s:Delta j⟶ X.  (((A ⟶ B))s (Delta ⊢ (A)s ⟶ (B)s) ∈ {Delta ⊢ _})

Lemma: cubical-fun-equal
[X:j⊢]. ∀[A,B:{X ⊢ _}]. ∀[f,g:{X ⊢ _:(A ⟶ B)}].
  g ∈ {X ⊢ _:(A ⟶ B)} 
  supposing ∀[I:fset(ℕ)]. ∀[a:X(I)]. ∀[J:fset(ℕ)]. ∀[h:J ⟶ I]. ∀[u:A(h(a))].  ((f(a) u) (g(a) u) ∈ B(h(a)))

Lemma: cubical-fun-equal2
[X:j⊢]. ∀[A,B:{X ⊢ _}]. ∀[f:{X ⊢ _:(A ⟶ B)}].
[g:I:fset(ℕ) ⟶ a:X(I) ⟶ J:fset(ℕ) ⟶ h:J ⟶ I ⟶ u:A(h(a)) ⟶ B(h(a))].
  g ∈ {X ⊢ _:(A ⟶ B)} 
  supposing ∀[I:fset(ℕ)]. ∀[a:X(I)]. ∀[J:fset(ℕ)]. ∀[h:J ⟶ I]. ∀[u:A(h(a))].  ((f(a) u) (g(a) u) ∈ B(h(a)))

Lemma: cubical-fun-p
X:j⊢. ∀A,B,T:{X ⊢ _}.  (((A ⟶ B))p (X.T ⊢ (A)p ⟶ (B)p) ∈ {X.T ⊢ _})

Lemma: member-cubical-fun-p
[X:j⊢]. ∀[A,B,T:{X ⊢ _}]. ∀[x:{X ⊢ _:(A ⟶ B)}].  ((x)p ∈ {X.T ⊢ _:((A)p ⟶ (B)p)})

Lemma: cc-snd_wf-cubical-fun
X:j⊢. ∀A,B:{X ⊢ _}.  (q ∈ {X.(A ⟶ B) ⊢ _:((A)p ⟶ (B)p)})

Lemma: cubical-fun-as-cubical-pi
[X:j⊢]. ∀[A,B:{X ⊢ _}].  ((X ⊢ A ⟶ B) X ⊢ Π(B)p ∈ {X ⊢ _})

Definition: cubical-lambda
b) ==  λI,a,J,f,u. (b (f(a);u))

Lemma: cubical-lambda_wf
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[B:{X.A ⊢ _}]. ∀[b:{X.A ⊢ _:B}].  ((λb) ∈ {X ⊢ _:ΠB})

Lemma: csm-cubical-lambda
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[B:{X.A ⊢ _}]. ∀[b:{X.A ⊢ _:B}]. ∀[H:j⊢]. ∀[s:H j⟶ X].  (((λb))s (b)s+) ∈ {H ⊢ _:(ΠB)s})

Definition: cubical-lam
cubical-lam(X;b) ==  b)

Lemma: cubical-lam_wf
[X:j⊢]. ∀[A,B:{X ⊢ _}]. ∀[b:{X.A ⊢ _:(B)p}].  (cubical-lam(X;b) ∈ {X ⊢ _:(A ⟶ B)})

Lemma: csm-cubical-lam
[X:j⊢]. ∀[A,B:{X ⊢ _}]. ∀[b:{X.A ⊢ _:(B)p}]. ∀[H:j⊢]. ∀[s:H j⟶ X].
  ((cubical-lam(X;b))s cubical-lam(H;(b)s+) ∈ {H ⊢ _:((A ⟶ B))s})

Definition: cubical-id-fun
cubical-id-fun(X) ==  cubical-lam(X;q)

Lemma: cubical-id-fun_wf
[X:j⊢]. ∀[A:{X ⊢ _}].  (cubical-id-fun(X) ∈ {X ⊢ _:(A ⟶ A)})

Lemma: csm-cubical-id-fun
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[H:j⊢]. ∀[s:H j⟶ X].  ((cubical-id-fun(X))s cubical-id-fun(H) ∈ {H ⊢ _:((A)s ⟶ (A)s)})

Definition: cubical-app
app(w; u) ==  λI,a. (w (u a))

Lemma: cubical-app_wf
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[B:{X.A ⊢ _}]. ∀[w:{X ⊢ _:ΠB}]. ∀[u:{X ⊢ _:A}].  (app(w; u) ∈ {X ⊢ _:(B)[u]})

Lemma: cubical-app_wf-csm
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[B:{X.A ⊢ _}]. ∀[w:{X ⊢ _:ΠB}]. ∀[H:j⊢]. ∀[tau:H j⟶ X]. ∀[u:{H ⊢ _:(A)tau}].
  (app((w)tau; u) ∈ {H ⊢ _:((B)tau+)[u]})

Definition: cubical-apply
cubical-apply(w;u) ==  app(w; u)

Lemma: cubical-apply_wf
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[B:{X.A ⊢ _}]. ∀[w:{X ⊢ _:ΠB}]. ∀[u:{X ⊢ _:A}].  (cubical-apply(w;u) ∈ {X ⊢ _:(B)[u]})

Lemma: cubical-app_wf_fun
[X:j⊢]. ∀[A,B:{X ⊢ _}]. ∀[w:{X ⊢ _:(A ⟶ B)}]. ∀[u:{X ⊢ _:A}].  (app(w; u) ∈ {X ⊢ _:B})

Definition: cubical-fun-comp
(f g) ==  cubical-lam(G;app((f)p; app((g)p; q)))

Lemma: cubical-fun-comp_wf
[X:j⊢]. ∀[A,B,C:{X ⊢ _}]. ∀[g:{X ⊢ _:(A ⟶ B)}]. ∀[f:{X ⊢ _:(B ⟶ C)}].  ((f g) ∈ {X ⊢ _:(A ⟶ C)})

Definition: cubical-sigma
Σ ==  <λI,a. (u:A(a) × B((a;u))), λI,J,f,a,p. <(fst(p) f), (snd(p) (a;fst(p)) f)>>

Lemma: cubical-sigma_wf
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[B:{X.A ⊢ _}].  X ⊢ Σ B

Lemma: cubical-sigma-at
[X,A,B,I,a:Top].  (Σ B(a) u:A(a) × B((a;u)))

Lemma: csm-cubical-sigma
X,Delta:j⊢. ∀A:{X ⊢ _}. ∀B:{X.A ⊢ _}. ∀s:Delta j⟶ X.  ((Σ B)s = Σ (A)s (B)(s p;q) ∈ {Delta ⊢ _})

Lemma: cubical-sigma-p
X:j⊢. ∀T,A:{X ⊢ _}. ∀B:{X.A ⊢ _}.  ((Σ B)p = Σ (A)p (B)(p p;q) ∈ {X.T ⊢ _})

Lemma: cubical-sigma-p-p
X:j⊢. ∀T,A:{X ⊢ _}. ∀B:{X.A ⊢ _}. ∀C:{X.T ⊢ _}.  (((Σ B)p)p = Σ ((A)p)p (B)(p p;q) ∈ {X.T.C ⊢ _})

Lemma: csm-cubical-sigma-typed
X,Delta:j⊢. ∀A:{X ⊢ _}. ∀B:{X.A ⊢ _}. ∀s:Delta j⟶ X.  ((Σ B)s = Σ (A)s (B)(s)dep ∈ {Delta ⊢ _})

Definition: cubical-fst
p.1 ==  λI,a. (fst((p a)))

Lemma: cubical-fst_wf
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[B:{X.A ⊢ _}]. ∀[p:{X ⊢ _:Σ B}].  (p.1 ∈ {X ⊢ _:A})

Lemma: csm-cubical-fst
[p,s:Top].  ((p.1)s (p)s.1)

Lemma: csm-ap-cubical-fst
[X,Delta:j⊢]. ∀[A:{X ⊢ _}]. ∀[B:{X.A ⊢ _}]. ∀[p:{X ⊢ _:Σ B}]. ∀[s:Delta j⟶ X].  ((p.1)s (p)s.1 ∈ {Delta ⊢ _:(A)s})

Definition: cubical-snd
p.2 ==  λI,a. (snd((p a)))

Lemma: cubical-snd_wf
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[B:{X.A ⊢ _}]. ∀[p:{X ⊢ _:Σ B}].  (p.2 ∈ {X ⊢ _:(B)[p.1]})

Lemma: cubical-snd-at
[a,I,p:Top].  (p.2(a) snd(p(a)))

Lemma: cubical-fst-at
[a,I,p:Top].  (p.1(a) fst(p(a)))

Lemma: csm-cubical-snd
[p,s:Top].  ((p.2)s (p)s.2)

Lemma: csm-ap-cubical-snd
[X,Delta:j⊢]. ∀[A:{X ⊢ _}]. ∀[B:{X.A ⊢ _}]. ∀[p:{X ⊢ _:Σ B}]. ∀[s:Delta j⟶ X].
  ((p.2)s (p)s.2 ∈ {Delta ⊢ _:((B)[p.1])s})

Definition: cubical-pair
cubical-pair(u;v) ==  λI,a. <a, a>

Lemma: cubical-pair_wf
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[B:{X.A ⊢ _}]. ∀[u:{X ⊢ _:A}]. ∀[v:{X ⊢ _:(B)[u]}].  (cubical-pair(u;v) ∈ {X ⊢ _:Σ B})

Lemma: csm-cubical-pair
[u,v,s:Top].  ((cubical-pair(u;v))s cubical-pair((u)s;(v)s))

Definition: sigma-elim-csm
SigmaElim ==  λI,x. let a,u,v in ((a;u);v)

Lemma: sigma-elim-csm_wf
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[B:{X.A ⊢ _}].  (SigmaElim ∈ X.Σ ij⟶ X.A.B)

Definition: sigma-unelim-csm
SigmaUnElim ==  λI,x. let au,v in let a,u au in (a;(u;v))

Lemma: sigma-unelim-csm_wf
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[B:{X.A ⊢ _}].  (SigmaUnElim ∈ X.A.B ij⟶ X.Σ B)

Lemma: sigma-elim-unelim
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[B:{X.A ⊢ _}].  (SigmaUnElim SigmaElim 1(X.Σ B) ∈ X.Σ ij⟶ X.Σ B)

Lemma: sigma-unelim-elim
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[B:{X.A ⊢ _}].  (SigmaElim SigmaUnElim 1(X.A.B) ∈ X.A.B ij⟶ X.A.B)

Lemma: sigma-unelim-elim-type
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[B:{X.A ⊢ _}]. ∀[T:{X.Σ B ⊢ _}].  (((T)SigmaUnElim)SigmaElim T ∈ {X.Σ B ⊢ _})

Lemma: sigma-unelim-elim-term
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[B:{X.A ⊢ _}]. ∀[T:{X.Σ B ⊢ _}]. ∀[t:{X.Σ B ⊢ _:T}].
  (((t)SigmaUnElim)SigmaElim t ∈ {X.Σ B ⊢ _:T})

Lemma: sigma-unelim-p
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[B:{X.A ⊢ _}].  (p SigmaUnElim p ∈ X.A.B ij⟶ X)

Lemma: sigma-unelim-p-type
[T:Top]. (((T)p)SigmaUnElim ((T)p)p)

Lemma: sigma-unelim-p-term
[t:Top]. (((t)p)SigmaUnElim ((t)p)p)

Lemma: sigma-elim-rule
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[B:{X.A ⊢ _}]. ∀[T:{X.Σ B ⊢ _}]. ∀[t:{X.A.B ⊢ _:(T)SigmaUnElim}].
  ((t)SigmaElim ∈ {X.Σ B ⊢ _:T})

Lemma: sigma-elim-equality-rule
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[B:{X.A ⊢ _}]. ∀[T:{X.Σ B ⊢ _}]. ∀[t1,t2:{X.A.B ⊢ _:(T)SigmaUnElim}].
  (t1)SigmaElim (t2)SigmaElim ∈ {X.Σ B ⊢ _:T} supposing t1 t2 ∈ {X.A.B ⊢ _:(T)SigmaUnElim}

Lemma: sigma-elim-equality-rule2
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[B:{X.A ⊢ _}]. ∀[T:{X.Σ B ⊢ _}]. ∀[t1:{X.A.B ⊢ _:(T)SigmaUnElim}]. ∀[t2:{X.Σ B ⊢ _:T}].
  (t1)SigmaElim t2 ∈ {X.Σ B ⊢ _:T} supposing t1 (t2)SigmaUnElim ∈ {X.A.B ⊢ _:(T)SigmaUnElim}

Lemma: csm-ap-cubical-pair
[X,Delta:j⊢]. ∀[A:{X ⊢ _}]. ∀[B:{X.A ⊢ _}]. ∀[u:{X ⊢ _:A}]. ∀[v:{X ⊢ _:(B)[u]}]. ∀[s:Delta j⟶ X].
  ((cubical-pair(u;v))s cubical-pair((u)s;(v)s) ∈ {Delta ⊢ _:(Σ B)s})

Lemma: csm-cubical-app
[w,u,s:Top].  ((app(w; u))s app((w)s; (u)s))

Lemma: csm-ap-cubical-app
[X,Delta:j⊢]. ∀[A:{X ⊢ _}]. ∀[B:{X.A ⊢ _}]. ∀[w:{X ⊢ _:ΠB}]. ∀[u:{X ⊢ _:A}]. ∀[s:Delta j⟶ X].
  ((app(w; u))s app((w)s; (u)s) ∈ {Delta ⊢ _:((B)[u])s})

Lemma: csm-ap-cubical-app-fun
[X,Delta:j⊢]. ∀[A,B:{X ⊢ _}]. ∀[w:{X ⊢ _:(A ⟶ B)}]. ∀[u:{X ⊢ _:A}]. ∀[s:Delta j⟶ X].
  ((app(w; u))s app((w)s; (u)s) ∈ {Delta ⊢ _:(B)s})

Lemma: csm-ap-cubical-lambda
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[B:{X.A ⊢ _}]. ∀[b:{X.A ⊢ _:B}]. ∀[H:j⊢]. ∀[s:H j⟶ X].  (((λb))s (b)s+) ∈ {H ⊢ _:(ΠB)s})

Lemma: cubical-snd-pair
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[B:{X.A ⊢ _}]. ∀[u:{X ⊢ _:A}]. ∀[v:{X ⊢ _:(B)[u]}].  (cubical-pair(u;v).2 v ∈ {X ⊢ _:(B)[u]})

Lemma: cubical-fst-pair
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[u:{X ⊢ _:A}]. ∀[v:Top].  (cubical-pair(u;v).1 u ∈ {X ⊢ _:A})

Lemma: cubical-pair-eta
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[B:{X.A ⊢ _}]. ∀[w:{X ⊢ _:Σ B}].  (cubical-pair(w.1;w.2) w ∈ {X ⊢ _:Σ B})

Lemma: cubical-sigma-equal
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[B:{X.A ⊢ _}]. ∀[w,y:{X ⊢ _:Σ B}].
  (w y ∈ {X ⊢ _:Σ B}) supposing ((w.2 y.2 ∈ {X ⊢ _:(B)[w.1]}) and (w.1 y.1 ∈ {X ⊢ _:A}))

Lemma: cubical-eta
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[B:{X.A ⊢ _}]. ∀[w:{X ⊢ _:ΠB}].  ((λapp((w)p; q)) w ∈ {X ⊢ _:ΠB})

Lemma: cubical-fun-eta
[X:j⊢]. ∀[A,B:{X ⊢ _}]. ∀[w:{X ⊢ _:(A ⟶ B)}].  (cubical-lam(X;app((w)p; q)) w ∈ {X ⊢ _:(A ⟶ B)})

Lemma: cubical-beta
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[B:{X.A ⊢ _}]. ∀[b:{X.A ⊢ _:B}]. ∀[u:{X ⊢ _:A}].  (app((λb); u) (b)[u] ∈ {X ⊢ _:(B)[u]})

Lemma: cubical-app-id-fun
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[u:{X ⊢ _:A}].  (app(cubical-id-fun(X); u) u ∈ {X ⊢ _:A})

Definition: cubical-sigma-fun
cubical-sigma-fun(G;A;B;f) ==  cubical-lam(G;cubical-pair(q.1;app(app(((λf))p; q.1); q.2)))

Lemma: cubical-sigma-fun_wf
[G:j⊢]. ∀[A:{G ⊢ _}]. ∀[B,B':{G.A ⊢ _}]. ∀[f:{G.A ⊢ _:(B ⟶ B')}].
  (cubical-sigma-fun(G;A;B;f) ∈ {G ⊢ _:(Σ B ⟶ Σ B')})

Definition: discrete-cubical-type
discr(T) ==  <λI,alpha. T, λI,J,f,alpha,x. x>

Lemma: discrete-cubical-type_wf
[T:Type]. ∀[X:j⊢].  X ⊢ discr(T)

Lemma: csm-discrete-cubical-type
[T,s:Top].  ((discr(T))s discr(T))

Definition: discrete-cubical-term
discr(t) ==  λI,alpha. t

Lemma: discrete-cubical-term_wf
[T:Type]. ∀[t:T]. ∀[X:j⊢].  (discr(t) ∈ {X ⊢ _:discr(T)})

Lemma: discrete-cubical-term-at-morph
[T:Type]. ∀[X:j⊢]. ∀[t:{X ⊢ _:discr(T)}].  ∀I,J:fset(ℕ). ∀f:J ⟶ I. ∀a:X(I).  (t(a) t(f(a)) ∈ T)

Lemma: discrete-unary_wf
[A,B:Type]. ∀[f:A ⟶ B]. ∀[X:j⊢]. ∀[t:{X ⊢ _:discr(A)}].  (discrete-unary(t;x.f[x]) ∈ {X ⊢ _:discr(B)})

Lemma: csm-discrete-cubical-term
[t,s:Top].  ((discr(t))s discr(t))

Lemma: discrete-cubical-term-is-map
[T:Type]. ∀[X:j⊢].  {X ⊢ _:discr(T)} ≡ ij⟶ discrete-cube(T)

Lemma: discrete-map-is-constant
[T:𝕌{j}]. ∀[I:fset(ℕ)]. ∀[s:formal-cube(I) ij⟶ discrete-cube(T)].
  (s J,g. (s 1)) ∈ formal-cube(I) ij⟶ discrete-cube(T))

Lemma: discrete-cubical-term-is-constant
Not every term of discrete cubical type is constant, but when the
context is the formal-cube(I) -- Yoneda(I) -- then it is.⋅

[T:Type]. ∀[I:fset(ℕ)]. ∀[t:{formal-cube(I) ⊢ _:discr(T)}].  (t discr(t(1)) ∈ {formal-cube(I) ⊢ _:discr(T)})

Lemma: discrete-map-is-constant2
[T:Type]. ∀[I:fset(ℕ)]. ∀[phi:𝔽(I)]. ∀[s:I,phi ⟶ discrete-cube(T)]. ∀[f:{f:I ⟶ I| (phi f) 1} ].
  J,g. (s f)) ∈ I,phi ⟶ discrete-cube(T) 
  supposing ∀[J:fset(ℕ)]. ∀[g:J ⟶ I].  ((phi g)  (g f ⋅ g ∈ J ⟶ I))

Lemma: discrete-cubical-term-is-constant-on-irr-face
Not every term of discrete cubical type is constant, but when the
context is the formal-cube(I) -- Yoneda(I) -- then it is.⋅

[T:Type]. ∀[I:fset(ℕ)]. ∀[as,bs:fset(names(I))].
  ∀[t:{I,irr_face(I;as;bs) ⊢ _:discr(T)}]. (t discr(t(irr-face-morph(I;as;bs))) ∈ {I,irr_face(I;as;bs) ⊢ _:discr(T)}) 
  supposing ↑fset-disjoint(NamesDeq;as;bs)

Definition: cubical-unit
==  discr(Unit)

Lemma: cubical-unit_wf
[X:j⊢]. X ⊢ 1

Lemma: csm-cubical-unit
[s:Top]. ((1)s 1)

Definition: cubical-it
==  discr(⋅)

Lemma: cubical-it_wf
[X:j⊢]. (* ∈ {X ⊢ _:1})

Lemma: cubical-it-unique
[X:j⊢]. ∀[t:{X ⊢ _:1}].  (t * ∈ {X ⊢ _:1})

Definition: constant-cubical-type
(X) ==  <λI,alpha. X(I), λI,J,f,alpha,w. f(w)>

Lemma: constant-cubical-type_wf
[X:SmallCubicalSet]. ∀[Gamma:j⊢].  Gamma ⊢ (X)

Lemma: csm-constant-cubical-type
[X,s:Top].  (((X))s (X))

Lemma: constant-cubical-type-at
[X,I,a:Top].  ((X)(a) X(I))

Lemma: unit-cube-cubical-type
[I:fset(ℕ)]
  ({formal-cube(I) ⊢ _} {AF:A:L:fset(ℕ) ⟶ L ⟶ I ⟶ Type × (L:fset(ℕ)
                                                              ⟶ J:fset(ℕ)
                                                              ⟶ f:J ⟶ L
                                                              ⟶ a:L ⟶ I
                                                              ⟶ (A a)
                                                              ⟶ (A a ⋅ f))| 
                           let A,F AF 
                           in (∀K:fset(ℕ). ∀a:K ⟶ I. ∀u:A a.  ((F u) u ∈ (A a)))
                              ∧ (∀L,J,K:fset(ℕ). ∀f:J ⟶ L. ∀g:K ⟶ J. ∀a:L ⟶ I. ∀u:A a.
                                   ((F f ⋅ u) (F a ⋅ (F u)) ∈ (A a ⋅ f ⋅ g)))} )

Lemma: empty-cubical-subset-I_cube
[I,J:fset(ℕ)].  I,0(J))

Lemma: empty-cubical-subset
[I:fset(ℕ)]. ∀[A,B:Top].  I,0 ⊢ <A, B>

Lemma: empty-cubical-subset-term
[I:fset(ℕ)]. ∀[phi:Point(face_lattice(I))].
  ∀[X,A,B:Top].  (A B ∈ {I,phi ⊢ _:X}) supposing phi 0 ∈ Point(face_lattice(I))

Lemma: member-empty-cubical-subset
[I:fset(ℕ)]. ∀[phi:Point(face_lattice(I))].
  ∀[X,A:Top].  (A ∈ {I,phi ⊢ _:X}) supposing phi 0 ∈ Point(face_lattice(I))

Definition: section-iota
section-iota(Gamma;A;I;rho;a) ==  (canonical-section(Gamma;A;I;rho;a))iota

Lemma: section-iota_wf
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[I:fset(ℕ)]. ∀[rho:Gamma(I)]. ∀[a:A(rho)]. ∀[psi:𝔽(I)].
  (section-iota(Gamma;A;I;rho;a) ∈ {I,psi ⊢ _:((A)<rho>)iota})

Lemma: cubical-term-at-comp
Gamma:j⊢. ∀T:{Gamma ⊢ _}. ∀t:{Gamma ⊢ _:T}. ∀I:fset(ℕ). ∀rho:Gamma(I). ∀J:fset(ℕ). ∀f:J ⟶ I. ∀K:fset(ℕ). ∀g:K ⟶ J.
  (t(f ⋅ g(rho)) t(g(f(rho))) ∈ T(g(f(rho))))

Lemma: subset-cubical-type-general
[X,Y:j⊢].  {X ⊢_} ⊆{Y ⊢_} supposing sub_cubical_set{j:l}(Y; X)

Lemma: subset-cubical-type
[X,Y:j⊢].  {X ⊢ _} ⊆{Y ⊢ _} supposing sub_cubical_set{j:l}(Y; X)

Lemma: subset-cubical-term
[X,Y:j⊢].  ∀[A:{X ⊢ _}]. ({X ⊢ _:A} ⊆{Y ⊢ _:A}) supposing sub_cubical_set{j:l}(Y; X)

Lemma: subset-cubical-term2
[X,Y:j⊢].  ∀[A,B:{X ⊢ _}].  {X ⊢ _:A} ⊆{Y ⊢ _:B} supposing B ∈ {X ⊢ _} supposing sub_cubical_set{j:l}(Y; X)

Lemma: csm-subset-codomain
[X,Y,Z:j⊢].  j⟶ Y ⊆j⟶ supposing sub_cubical_set{j:l}(Y; Z)

Lemma: csm-subset-domain
[X,Y,Z:j⊢].  j⟶ X ⊆j⟶ supposing sub_cubical_set{j:l}(Z; Y)

Lemma: csm-subset-subtype
[A,B,Y,Z:j⊢].  (Y j⟶ A ⊆j⟶ B) supposing (sub_cubical_set{j:l}(A; B) and sub_cubical_set{j:l}(Z; Y))

Lemma: sub_cubical_set_functionality
[Y,X:j⊢]. ∀[A:{X ⊢ _}].  sub_cubical_set{[i j]:l}(Y.A; X.A) supposing sub_cubical_set{j:l}(Y; X)

Lemma: cubical-pi-intro
G:j⊢. ∀A:{G ⊢ _}. ∀B:{G.A ⊢ _}.  ({G.A ⊢ _:B}  {G ⊢ _:ΠB})

Lemma: cubical-sigma-intro
G:j⊢. ∀A:{G ⊢ _}. ∀B:{G.A ⊢ _}.  ((∃a:{G ⊢ _:A}. {G ⊢ _:(B)[a]})  {G ⊢ _:Σ B})

Lemma: cubical-pi-implies-sigma
G:j⊢. ∀A:{G ⊢ _}. ∀B:{G.A ⊢ _}.  ({G ⊢ _:A}  {G ⊢ _:ΠB}  {G ⊢ _:Σ B})

Lemma: csm_typedp_id_adjoin_lemma
E,s,F,G,u,H:Top.  ((((E)s)tp{i:l})[u] (E)s)

Lemma: cubical-fun-family-discrete
[A,B:Type]. ∀[X:j⊢]. ∀[I:fset(ℕ)]. ∀[a:X(I)].
  (cubical-fun-family(X; discr(A); discr(B); I; a)
  {w:J:fset(ℕ) ⟶ f:J ⟶ I ⟶ u:A ⟶ B| ∀J,K:fset(ℕ). ∀f:J ⟶ I. ∀g:K ⟶ J. ∀u:A.  ((w u) (w f ⋅ u) ∈ B)} 
  ∈ Type)

Definition: dcff-inj
dcff-inj(I;w) ==  1

Lemma: dcff-inj_wf
[A,B:Type]. ∀[X:j⊢]. ∀[I:fset(ℕ)]. ∀[a:X(I)]. ∀[w:cubical-fun-family(X; discr(A); discr(B); I; a)].
  (dcff-inj(I;w) ∈ A ⟶ B)

Lemma: dcff-inj-injection
[A,B:Type]. ∀[X:j⊢]. ∀[I:fset(ℕ)]. ∀[a:X(I)].
  Inj(cubical-fun-family(X; discr(A); discr(B); I; a);A ⟶ B;λw.dcff-inj(I;w))

Lemma: dcff-inj-bijection
[A,B:Type]. ∀[X:j⊢]. ∀[I:fset(ℕ)]. ∀[a:X(I)].
  Bij(cubical-fun-family(X; discr(A); discr(B); I; a);A ⟶ B;λw.dcff-inj(I;w))

Comment: interval type doc
================================= interval type  ===============================
 From the interval presheaf 𝕀 we get cubical type 𝕀 and operations on it.⋅

Definition: interval-type
𝕀 ==  (𝕀)

Lemma: interval-type_wf
[Gamma:j⊢]. Gamma ⊢ 𝕀

Lemma: interval-type-ap-morph
[I,J,f,rho,u:Top].  ((u rho f) dM-lift(J;I;f) u)

Lemma: interval-type-ap-inc
[I,J:fset(ℕ)]. ∀[f:I ⟶ J]. ∀[x:names(J)]. ∀[rho:Top].  ((<x> rho f) (f x) ∈ Point(dM(I)))

Lemma: interval-type-at
[J,rho:Top].  (𝕀(rho) ~ 𝕀(J))

Lemma: interval-type-at-is-point
[J,rho:Top].  (𝕀(rho) Point(dM(J)))

Lemma: csm-interval-type
[s:Top]. ((𝕀)s ~ 𝕀)

Lemma: csm+-comp-csm+-sq-interval
[H,K,X,tau,s:Top].  (tau+ s+ tau s+)

Definition: interval-0
0(𝕀==  λI,rho. 0

Lemma: interval-0_wf
[Gamma:j⊢]. (0(𝕀) ∈ {Gamma ⊢ _:𝕀})

Lemma: csm-id-adjoin_wf-interval-0
[Gamma:j⊢]. ([0(𝕀)] ∈ Gamma j⟶ Gamma.𝕀)

Lemma: csm-interval-0
[s:Top]. ((0(𝕀))s 0(𝕀))

Lemma: cc-fst_wf_interval
[Gamma:j⊢]. (p ∈ Gamma.𝕀 j⟶ Gamma)

Lemma: cc-snd-0
[X:Top]. ((q)[0(𝕀)] 0(𝕀))

Lemma: interval-0-at
[I,a:Top].  (0(𝕀)(a) 0)

Definition: interval-1
1(𝕀==  λI,rho. 1

Lemma: interval-1_wf
[Gamma:j⊢]. (1(𝕀) ∈ {Gamma ⊢ _:𝕀})

Lemma: csm-id-adjoin_wf-interval-1
[Gamma:j⊢]. ([1(𝕀)] ∈ Gamma j⟶ Gamma.𝕀)

Lemma: csm-interval-1
[s:Top]. ((1(𝕀))s 1(𝕀))

Lemma: cc-snd-1
[X:Top]. ((q)[1(𝕀)] 1(𝕀))

Lemma: csm-ap-interval-1-adjoin-lemma
H:j⊢. ∀I:fset(ℕ). ∀v:H(I). ∀j:{i:ℕ| ¬i ∈ I} .  ((j1)((s(v);<j>)) ([1(𝕀)])v ∈ H.𝕀(I))

Lemma: interval-1-at
[I,a:Top].  (1(𝕀)(a) 1)

Definition: interval-rev
1-(r) ==  λI,rho. ¬(r(rho))

Lemma: interval-rev_wf
[Gamma:j⊢]. ∀[r:{Gamma ⊢ _:𝕀}].  (1-(r) ∈ {Gamma ⊢ _:𝕀})

Lemma: interval-rev-0
[Gamma:j⊢]. (1-(0(𝕀)) 1(𝕀) ∈ {Gamma ⊢ _:𝕀})

Lemma: interval-rev-1
[Gamma:j⊢]. (1-(1(𝕀)) 0(𝕀) ∈ {Gamma ⊢ _:𝕀})

Definition: interval-meet
r ∧ ==  λI,rho. r(rho) ∧ s(rho)

Lemma: interval-meet_wf
[Gamma:j⊢]. ∀[r,s:{Gamma ⊢ _:𝕀}].  (r ∧ s ∈ {Gamma ⊢ _:𝕀})

Lemma: csm-interval-meet
[h,r,s:Top].  ((r ∧ s)h (r)h ∧ (s)h)

Lemma: interval-meet-1
[X:j⊢]. ∀[z:{X ⊢ _:𝕀}].  (z ∧ 1(𝕀z ∈ {X ⊢ _:𝕀})

Lemma: interval-meet-0
[X:j⊢]. ∀[z:{X ⊢ _:𝕀}].  (z ∧ 0(𝕀0(𝕀) ∈ {X ⊢ _:𝕀})

Lemma: interval-meet-comm
[X:j⊢]. ∀[z,y:{X ⊢ _:𝕀}].  (z ∧ y ∧ z ∈ {X ⊢ _:𝕀})

Definition: interval-join
r ∨ ==  λI,rho. r(rho) ∨ s(rho)

Lemma: interval-join_wf
[Gamma:j⊢]. ∀[r,s:{Gamma ⊢ _:𝕀}].  (r ∨ s ∈ {Gamma ⊢ _:𝕀})

Lemma: csm-interval-join
[h,r,s:Top].  ((r ∨ s)h (r)h ∨ (s)h)

Lemma: csm+_wf_interval
[H,K:j⊢]. ∀[tau:K j⟶ H].  (tau+ ∈ K.𝕀 j⟶ H.𝕀)

Lemma: csm+-comp-csm+-interval
[H,K,X:j⊢]. ∀[tau:K j⟶ H]. ∀[s:X j⟶ K].  (tau+ s+ tau s+ ∈ X.𝕀 j⟶ H.𝕀)

Definition: cube+
cube+(I;i) ==  λA,fx. let f,x fx in λj.if (j =z i) then else fi 

Lemma: cube+_wf
[I:fset(ℕ)]. ∀[i:ℕ].  (cube+(I;i) ∈ formal-cube(I).𝕀 j⟶ formal-cube(I+i))

Definition: cube-
cube-(I;i) ==  λA,f. <f, i>

Lemma: cube-_wf
[I:fset(ℕ)]. ∀[i:ℕ].  (cube-(I;i) ∈ formal-cube(I+i) j⟶ formal-cube(I).𝕀)

Lemma: cube+-
[I:fset(ℕ)]. ∀[i:{i:ℕ| ¬i ∈ I} ].
  (cube-(I;i) cube+(I;i) 1(formal-cube(I).𝕀) ∈ formal-cube(I).𝕀 j⟶ formal-cube(I).𝕀)

Lemma: cube-+
[I:fset(ℕ)]. ∀[i:ℕ].  (cube+(I;i) cube-(I;i) 1(formal-cube(I+i)) ∈ formal-cube(I+i) j⟶ formal-cube(I+i))

Lemma: cube+_interval-0
[I:fset(ℕ)]. ∀[i:ℕ].  (cube+(I;i) [0(𝕀)] = <(i0)> ∈ formal-cube(I) j⟶ formal-cube(I+i))

Lemma: cube+_interval-1
[I:fset(ℕ)]. ∀[i:ℕ].  (cube+(I;i) [1(𝕀)] = <(i1)> ∈ formal-cube(I) j⟶ formal-cube(I+i))

Lemma: context-map-cube+-csm+
[Gamma:j⊢]. ∀[I,J:fset(ℕ)]. ∀[i:{i:ℕ| ¬i ∈ I} ]. ∀[j:{j:ℕ| ¬j ∈ J} ]. ∀[g:J ⟶ I]. ∀[rho:Gamma(I+i)].
  (<g,i=j(rho)> cube+(J;j) = <rho> cube+(I;i) o <g>+ ∈ formal-cube(J).𝕀 j⟶ Gamma)

Definition: swap-interval
swap-interval(G;A) ==  csm-swap(G;𝕀;A)

Lemma: swap-interval_wf
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}].  (swap-interval(Gamma;A) ∈ Gamma.𝕀.(A)p j⟶ Gamma.A.𝕀)

Lemma: p+-swap-interval
[Gamma:j⊢]. ∀[A:{Gamma.𝕀 ⊢ _}]. ∀[B:{Gamma ⊢ _}].  (((A)p+)swap-interval(Gamma;B) (A)p ∈ {Gamma.𝕀.(B)p ⊢ _})

Definition: swap-intervals
swap-intervals(G) ==  csm-swap(G;𝕀;𝕀)

Lemma: swap-intervals_wf
[Gamma:j⊢]. (swap-intervals(Gamma) ∈ Gamma.𝕀.𝕀 j⟶ Gamma.𝕀.𝕀)

Definition: example-bool-term
example-bool-term(X) ==  discrete-bool-term(I,alpha.isdM0(snd(λi.0(alpha))))

Comment: face type doc
=================================== face type  =================================

From the face presheaf 𝔽 we get cubcial type 𝔽 and operations on it.




Definition: face-type
𝔽 ==  (𝔽)

Lemma: face-type_wf
[Gamma:j⊢]. Gamma ⊢ 𝔽

Lemma: face-type-sq
𝔽 ~ <λI,alpha. Point(face_lattice(I)), λI,J,f,alpha,w. (w)<f>>

Lemma: face-type-ap-morph
[I,J,f,rho,u:Top].  ((u rho f) (u)<f>)

Lemma: face-type-ap-morph-ps
[I,J,f,rho,u:Top].  ((u rho f) (u)<f>)

Lemma: face-type-at
[J,rho:Top].  (𝔽(rho) Point(face_lattice(J)))

Lemma: cubical-type-at_wf_face-type
[J:fset(ℕ)]. ∀[rho:Top].  (𝔽(rho) ∈ Type)

Lemma: csm-face-type
[s:Top]. ((𝔽)s ~ 𝔽)

Definition: face-0
0(𝔽==  λI,rho. 0

Lemma: face-0_wf
[Gamma:j⊢]. (0(𝔽) ∈ {Gamma ⊢ _:𝔽})

Lemma: csm-face-0
[s:Top]. ((0(𝔽))s 0(𝔽))

Definition: face-1
1(𝔽==  λI,rho. 1

Lemma: face-1_wf
[Gamma:j⊢]. (1(𝔽) ∈ {Gamma ⊢ _:𝔽})

Lemma: csm-face-1
[s:Top]. ((1(𝔽))s 1(𝔽))

Definition: face-and
(a ∧ b) ==  λI,rho. a(rho) ∧ b(rho)

Lemma: face-and_wf
[Gamma:j⊢]. ∀[r,s:{Gamma ⊢ _:𝔽}].  ((r ∧ s) ∈ {Gamma ⊢ _:𝔽})

Lemma: face-and-com
[Gamma:j⊢]. ∀[r,s:{Gamma ⊢ _:𝔽}].  ((r ∧ s) (s ∧ r) ∈ {Gamma ⊢ _:𝔽})

Lemma: face-and-at
[r,s,I,rho:Top].  ((r ∧ s)(rho) r(rho) ∧ s(rho))

Lemma: face-and-eq-1
[Gamma:j⊢]
  ∀r,s:{Gamma ⊢ _:𝔽}. ∀I:fset(ℕ). ∀rho:Gamma(I).
    ((r ∧ s)(rho) 1 ∈ Point(face_lattice(I))
    ⇐⇒ (r(rho) 1 ∈ Point(face_lattice(I))) ∧ (s(rho) 1 ∈ Point(face_lattice(I))))

Definition: face-or
(a ∨ b) ==  λI,rho. a(rho) ∨ b(rho)

Lemma: face-or_wf
[Gamma:j⊢]. ∀[r,s:{Gamma ⊢ _:𝔽}].  ((r ∨ s) ∈ {Gamma ⊢ _:𝔽})

Lemma: face-or-at
[r,s,I,rho:Top].  ((r ∨ s)(rho) r(rho) ∨ s(rho))

Lemma: face-or-1
[X:j⊢]. ∀[psi:{X ⊢ _:𝔽}].  ((psi ∨ 1(𝔽)) 1(𝔽) ∈ {X ⊢ _:𝔽})

Lemma: face-or-0
[X:j⊢]. ∀[psi:{X ⊢ _:𝔽}].  ((psi ∨ 0(𝔽)) psi ∈ {X ⊢ _:𝔽})

Lemma: face-1-or
[X:j⊢]. ∀[psi:{X ⊢ _:𝔽}].  ((1(𝔽) ∨ psi) 1(𝔽) ∈ {X ⊢ _:𝔽})

Lemma: face-0-or
[X:j⊢]. ∀[psi:{X ⊢ _:𝔽}].  ((0(𝔽) ∨ psi) psi ∈ {X ⊢ _:𝔽})

Lemma: csm-face-or
[r,s,sigma:Top].  (((r ∨ s))sigma ((r)sigma ∨ (s)sigma))

Lemma: csm-face-and
[r,s,sigma:Top].  (((r ∧ s))sigma ((r)sigma ∧ (s)sigma))

Lemma: face-or-eq-1
[Gamma:j⊢]
  ∀r,s:{Gamma ⊢ _:𝔽}. ∀I:fset(ℕ). ∀rho:Gamma(I).
    ((r ∨ s)(rho) 1 ∈ Point(face_lattice(I))
    ⇐⇒ (r(rho) 1 ∈ Point(face_lattice(I))) ∨ (s(rho) 1 ∈ Point(face_lattice(I))))

Definition: face-or-list
\/(L) ==  reduce(λa,b. (a ∨ b);0(𝔽);L)

Lemma: face-or-list_wf
[Gamma:j⊢]. ∀[L:{Gamma ⊢ _:𝔽List].  (\/(L) ∈ {Gamma ⊢ _:𝔽})

Lemma: face-or-list-eq-1
[Gamma:j⊢]
  ∀L:{Gamma ⊢ _:𝔽List. ∀I:fset(ℕ). ∀rho:Gamma(I).
    (\/(L)(rho) 1 ∈ Point(face_lattice(I)) ⇐⇒ (∃x∈L. x(rho) 1 ∈ Point(face_lattice(I))))

Definition: face-one
(i=1) ==  λI,rho. dM-to-FL(I;i(rho))

Lemma: face-one_wf
[Gamma:j⊢]. ∀[i:{Gamma ⊢ _:𝕀}].  ((i=1) ∈ {Gamma ⊢ _:𝔽})

Definition: face-zero
(i=0) ==  λI,rho. dM-to-FL(I;¬(i(rho)))

Lemma: face-zero_wf
[Gamma:j⊢]. ∀[i:{Gamma ⊢ _:𝕀}].  ((i=0) ∈ {Gamma ⊢ _:𝔽})

Lemma: face-zero-as-face-one
[Gamma:j⊢]. ∀[i:{Gamma ⊢ _:𝕀}].  ((i=0) (1-(i)=1) ∈ {Gamma ⊢ _:𝔽})

Lemma: csm-face-zero
[r,sigma:Top].  (((r=0))sigma ((r)sigma=0))

Lemma: csm-face-one
[r,sigma:Top].  (((r=1))sigma ((r)sigma=1))

Lemma: face-one-and-zero
[X:j⊢]. ∀[z:{X ⊢ _:𝕀}].  (((z=1) ∧ (z=0)) 0(𝔽) ∈ {X ⊢ _:𝔽})

Lemma: face-zero-and-one
[X:j⊢]. ∀[z:{X ⊢ _:𝕀}].  (((z=0) ∧ (z=1)) 0(𝔽) ∈ {X ⊢ _:𝔽})

Lemma: face-zero-eq-1
[H:j⊢]. ∀[z:{H ⊢ _:𝕀}]. ∀[I:fset(ℕ)]. ∀[a:H(I)].  z(a) 0 ∈ 𝕀(a) supposing (z=0)(a) 1 ∈ Point(face_lattice(I))

Lemma: face-zero-interval-0
[H:j⊢]. ((0(𝕀)=0) 1(𝔽) ∈ {H ⊢ _:𝔽})

Lemma: face-zero-interval-1
[H:j⊢]. ((1(𝕀)=0) 0(𝔽) ∈ {H ⊢ _:𝔽})

Lemma: face-one-interval-1
[H:j⊢]. ((1(𝕀)=1) 1(𝔽) ∈ {H ⊢ _:𝔽})

Lemma: face-one-interval-0
[H:j⊢]. ((0(𝕀)=1) 0(𝔽) ∈ {H ⊢ _:𝔽})

Lemma: face-one-eq-1
[H:j⊢]. ∀[z:{H ⊢ _:𝕀}]. ∀[I:fset(ℕ)]. ∀[a:H(I)].  z(a) 1 ∈ 𝕀(a) supposing (z=1)(a) 1 ∈ Point(face_lattice(I))

Lemma: face-term-at-restriction
[Gamma:j⊢]. ∀[phi:{Gamma ⊢ _:𝔽}]. ∀[I,J:fset(ℕ)]. ∀[f:J ⟶ I]. ∀[a:Gamma(I)].  (phi(f(a)) f(phi(a)) ∈ 𝔽(J))

Lemma: face-term-at-restriction-eq-1
[Gamma:j⊢]. ∀[phi:{Gamma ⊢ _:𝔽}]. ∀[A,B:fset(ℕ)]. ∀[g:B ⟶ A]. ∀[v:Gamma(A)].
  phi(g(v)) 1 ∈ Point(face_lattice(B)) supposing phi(v) 1 ∈ Point(face_lattice(A))

Definition: context-subset
Gamma, phi ==  <λI.{rho:Gamma(I)| phi(rho) 1 ∈ Point(face_lattice(I))} , λI,J,f,rho. f(rho)>

Lemma: context-subset_wf
[Gamma:j⊢]. ∀[phi:{Gamma ⊢ _:𝔽}].  Gamma, phi j⊢

Lemma: context-subset-restriction
[Gamma,phi,I,J,a,f:Top].  (f(a) f(a))

Lemma: context-subset-map
[Gamma:j⊢]. ∀[phi:{Gamma ⊢ _:𝔽}]. ∀[Z:j⊢]. ∀[s:Z j⟶ Gamma].  (s ∈ Z, (phi)s j⟶ Gamma, phi)

Lemma: context-map-subset
[G:j⊢]. ∀[phi:{G ⊢ _:𝔽}]. ∀[I:fset(ℕ)]. ∀[a:G, phi(I)].  (<a> = <a> ∈ formal-cube(I) j⟶ G, phi)

Lemma: cubical-fun-subset
[G,phi,T,A:Top].  ((G, phi ⊢ T ⟶ A) (G ⊢ T ⟶ A))

Lemma: cubical-fun-subset-adjoin
[G,B,phi,T,A:Top].  ((G, phi.B ⊢ T ⟶ A) (G.B ⊢ T ⟶ A))

Lemma: subset-iota_wf2
[Gamma:j⊢]. ∀[phi:{Gamma ⊢ _:𝔽}].  (iota ∈ Gamma, phi j⟶ Gamma)

Lemma: cubical-subset-is-context-subset
[I:fset(ℕ)]. ∀[psi:𝔽(I)].  (I,psi formal-cube(I), λJ,f. f(psi) ∈ CubicalSet{j})

Lemma: cubical-subset-is-context-subset-canonical
[I:fset(ℕ)]. ∀[psi:𝔽(I)].  (I,psi formal-cube(I), canonical-section(();𝔽;I;⋅;psi) ∈ CubicalSet{j})

Lemma: context-subset-is-cubical-subset
[I:fset(ℕ)]. ∀[phi:{formal-cube(I) ⊢ _:𝔽}].  (formal-cube(I), phi I,phi(1) ∈ CubicalSet{j})

Definition: face-term-implies
Gamma ⊢ (phi  psi) ==
  ∀I:fset(ℕ). ∀rho:Gamma(I).  ((phi(rho) 1 ∈ Point(face_lattice(I)))  (psi(rho) 1 ∈ Point(face_lattice(I))))

Lemma: face-term-implies_wf
[Gamma:j⊢]. ∀[phi,psi:{Gamma ⊢ _:𝔽}].  (Gamma ⊢ (phi  psi) ∈ ℙ{[i j']})

Lemma: csm-face-term-implies
[Gamma:j⊢]. ∀[phi,psi:{Gamma ⊢ _:𝔽}].
  ∀[H:j⊢]. ∀[s:H j⟶ Gamma].  H ⊢ ((phi)s  (psi)s) supposing Gamma ⊢ (phi  psi)

Lemma: face-term-implies-1
[Gamma:j⊢]. ∀[phi:{Gamma ⊢ _:𝔽}].  Gamma ⊢ (phi  1(𝔽))

Lemma: face-term-implies-same
[Gamma:j⊢]. ∀[phi:{Gamma ⊢ _:𝔽}].  Gamma ⊢ (phi  phi)

Lemma: face-term-and-implies1
[Gamma:j⊢]. ∀[phi,psi:{Gamma ⊢ _:𝔽}].  Gamma ⊢ ((phi ∧ psi)  phi)

Lemma: face-term-and-implies2
[Gamma:j⊢]. ∀[phi,psi:{Gamma ⊢ _:𝔽}].  Gamma ⊢ ((phi ∧ psi)  psi)

Lemma: face-term-and-implies
[Gamma:j⊢]. ∀[phi,psi,phi',psi':{Gamma ⊢ _:𝔽}].
  (Gamma ⊢ ((phi ∧ psi)  (phi' ∧ psi'))) supposing (Gamma ⊢ (phi  phi') and Gamma ⊢ (psi  psi'))

Lemma: face-term-implies-or1
[Gamma:j⊢]. ∀[phi,psi:{Gamma ⊢ _:𝔽}].  Gamma ⊢ (phi  (phi ∨ psi))

Lemma: face-term-implies-or2
[Gamma:j⊢]. ∀[phi,psi:{Gamma ⊢ _:𝔽}].  Gamma ⊢ (phi  (psi ∨ phi))

Lemma: face-term-or-implies
[Gamma:j⊢]. ∀[a,b,c:{Gamma ⊢ _:𝔽}].  (Gamma ⊢ ((a ∨ b)  c)) supposing (Gamma ⊢ (a  c) and Gamma ⊢ (b  c))

Lemma: face-term-implies-and
[Gamma:j⊢]. ∀[a,b,c:{Gamma ⊢ _:𝔽}].  (Gamma ⊢ (c  (a ∧ b))) supposing (Gamma ⊢ (c  a) and Gamma ⊢ (c  b))

Lemma: face-term-implies-or
[Gamma:j⊢]. ∀[a,b,c:{Gamma ⊢ _:𝔽}].  ((Gamma ⊢ (c  a) ∨ Gamma ⊢ (c  b))  Gamma ⊢ (c  (a ∨ b)))

Definition: face-term-iff
Gamma ⊢ (phi ⇐⇒ psi) ==  Gamma ⊢ (phi  psi) ∧ Gamma ⊢ (psi  phi)

Lemma: face-term-iff_wf
[Gamma:j⊢]. ∀[phi,psi:{Gamma ⊢ _:𝔽}].  (Gamma ⊢ (phi ⇐⇒ psi) ∈ ℙ{[i j']})

Lemma: face-term-implies_functionality
[Gamma:j⊢]. ∀[a,b,c,d:{Gamma ⊢ _:𝔽}].
  (Gamma ⊢ (a ⇐⇒ c)  Gamma ⊢ (b ⇐⇒ d)  (Gamma ⊢ (a  b) ⇐⇒ Gamma ⊢ (c  d)))

Lemma: face-term-distrib1
[Gamma:j⊢]. ∀[a,b,c:{Gamma ⊢ _:𝔽}].  Gamma ⊢ ((a ∧ (b ∨ c)) ⇐⇒ ((a ∧ b) ∨ (a ∧ c)))

Lemma: face-term-distrib2
[Gamma:j⊢]. ∀[a,b,c:{Gamma ⊢ _:𝔽}].  Gamma ⊢ ((a ∨ (b ∧ c)) ⇐⇒ ((a ∨ b) ∧ (a ∨ c)))

Lemma: face-term-distrib3
[Gamma:j⊢]. ∀[a,b,c:{Gamma ⊢ _:𝔽}].  Gamma ⊢ (((b ∨ c) ∧ a) ⇐⇒ ((b ∧ a) ∨ (c ∧ a)))

Lemma: face-term-distrib4
[Gamma:j⊢]. ∀[a,b,c:{Gamma ⊢ _:𝔽}].  Gamma ⊢ (((b ∧ c) ∨ a) ⇐⇒ ((b ∨ a) ∧ (c ∨ a)))

Definition: face-forall
(∀ phi) ==  λJ,rho. (∀new-name(J).phi((s(rho);<new-name(J)>)))

Lemma: face-forall_wf
[Gamma:j⊢]. ∀[phi:{Gamma.𝕀 ⊢ _:𝔽}].  ((∀ phi) ∈ {Gamma ⊢ _:𝔽})

Lemma: face-forall-1
[Gamma:j⊢]. ∀[phi:{Gamma.𝕀 ⊢ _:𝔽}].  (Gamma ⊢ ∀ phi) 1(𝔽) ∈ {Gamma ⊢ _:𝔽supposing phi 1(𝔽) ∈ {Gamma.𝕀 ⊢ _:𝔽}

Lemma: csm-face-forall
[Gamma,Delta:j⊢]. ∀[sigma:Delta j⟶ Gamma]. ∀[phi:{Gamma.𝕀 ⊢ _:𝔽}].
  (((∀ phi))sigma (Delta ⊢ ∀ (phi)sigma+) ∈ {Delta ⊢ _:𝔽})

Lemma: face-forall-or
[Gamma:j⊢]. ∀[phi,psi:{Gamma.𝕀 ⊢ _:𝔽}].  ((Gamma ⊢ ∀ (phi ∨ psi)) ((∀ phi) ∨ (∀ psi)) ∈ {Gamma ⊢ _:𝔽})

Lemma: face-forall-q=0
[Gamma:j⊢]. ((Gamma ⊢ ∀ (q=0)) 0(𝔽) ∈ {Gamma ⊢ _:𝔽})

Lemma: face-forall-q=1
[Gamma:j⊢]. ((Gamma ⊢ ∀ (q=1)) 0(𝔽) ∈ {Gamma ⊢ _:𝔽})

Lemma: face-forall-q=0-or-q=1
[Gamma:j⊢]. ((Gamma ⊢ ∀ ((q=0) ∨ (q=1))) 0(𝔽) ∈ {Gamma ⊢ _:𝔽})

Lemma: face-forall-and
[Gamma:j⊢]. ∀[phi,psi:{Gamma.𝕀 ⊢ _:𝔽}].  ((Gamma ⊢ ∀ (phi ∧ psi)) ((∀ phi) ∧ (∀ psi)) ∈ {Gamma ⊢ _:𝔽})

Lemma: csm-canonical-section-face
[X,I,rho,phi,s:Top].  ((canonical-section(X;𝔽;I;rho;phi))s ~ λJ,a. (phi)<(s)a>)

Lemma: csm-canonical-section-face-type
[I,K:fset(ℕ)]. ∀[f:K ⟶ I]. ∀[phi:𝔽(I)].
  (canonical-section(();𝔽;K;⋅;f(phi)) (canonical-section(();𝔽;I;⋅;phi))<f> ∈ {formal-cube(K) ⊢ _:𝔽})

Lemma: csm-canonical-section-face-type-0
[I:fset(ℕ)]. ∀[i:{i:ℕ| ¬i ∈ I} ]. ∀[phi:𝔽(I)].
  ((canonical-section(();𝔽;I+i;⋅;s(phi)))<(i0)> canonical-section(();𝔽;I;⋅;phi) ∈ {formal-cube(I) ⊢ _:𝔽})

Lemma: csm-canonical-section-face-type-1
[I:fset(ℕ)]. ∀[i:{i:ℕ| ¬i ∈ I} ]. ∀[phi:𝔽(I)].
  ((canonical-section(();𝔽;I+i;⋅;s(phi)))<(i1)> canonical-section(();𝔽;I;⋅;phi) ∈ {formal-cube(I) ⊢ _:𝔽})

Lemma: context-subset-subtype
[Gamma:j⊢]. ∀[phi1,phi2:{Gamma ⊢ _:𝔽}].
  {Gamma, phi1 ⊢ _} ⊆{Gamma, phi2 ⊢ _} 
  supposing ∀I:fset(ℕ). ∀rho:Gamma(I).
              ((phi2(rho) 1 ∈ Point(face_lattice(I)))  (phi1(rho) 1 ∈ Point(face_lattice(I))))

Lemma: context-subset_functionality
[Gamma:j⊢]. ∀[a,b:{Gamma ⊢ _:𝔽}].  (Gamma ⊢ (a ⇐⇒ b)  {Gamma, a ⊢ _} ≡ {Gamma, b ⊢ _})

Lemma: context-subset-is-subset
[Gamma:j⊢]. ∀[phi:{Gamma ⊢ _:𝔽}].  sub_cubical_set{j:l}(Gamma, phi; Gamma)

Lemma: cubical-term-restriction-is-1
[Gamma:j⊢]. ∀[phi:{Gamma ⊢ _:𝔽}]. ∀[I:fset(ℕ)]. ∀[rho:Gamma(I)]. ∀[J:fset(ℕ)]. ∀[f:J ⟶ I].
  ((phi(rho) 1 ∈ Point(face_lattice(I)))  (phi(f(rho)) 1 ∈ Point(face_lattice(J))))

Lemma: face-term-implies-subset
[Gamma:j⊢]. ∀[phi,psi:{Gamma ⊢ _:𝔽}].  sub_cubical_set{j:l}(Gamma, phi; Gamma, psi) supposing Gamma ⊢ (phi  psi)

Lemma: cubical-subset_functionality_wrt_le
J:fset(ℕ). ∀a,b:Point(face_lattice(J)).  (a ≤  sub_cubical_set{j:l}(J,a; J,b))

Lemma: context-subset-subtype-and
[Gamma:j⊢]. ∀[phi1,phi2:{Gamma ⊢ _:𝔽}].  ({Gamma, phi1 ⊢ _} ⊆{Gamma, (phi1 ∧ phi2) ⊢ _})

Lemma: context-subset-subtype-and2
[Gamma:j⊢]. ∀[phi1,phi2:{Gamma ⊢ _:𝔽}].  ({Gamma, phi2 ⊢ _} ⊆{Gamma, (phi1 ∧ phi2) ⊢ _})

Lemma: context-subset-subtype-or
[Gamma:j⊢]. ∀[phi1,phi2:{Gamma ⊢ _:𝔽}].  ({Gamma, (phi1 ∨ phi2) ⊢ _} ⊆{Gamma, phi1 ⊢ _})

Lemma: context-subset-subtype-or2
[Gamma:j⊢]. ∀[phi1,phi2:{Gamma ⊢ _:𝔽}].  ({Gamma, (phi1 ∨ phi2) ⊢ _} ⊆{Gamma, phi2 ⊢ _})

Lemma: cubical-term-at-comp-constant-type
[Gamma:j⊢]. ∀[phi:{Gamma ⊢ _:𝔽}]. ∀[I:fset(ℕ)]. ∀[rho:Gamma(I)]. ∀[J:fset(ℕ)]. ∀[f:{f:J ⟶ I| 
                                                                                     phi(f(rho))
                                                                                     1
                                                                                     ∈ Point(face_lattice(J))} ].
[K:fset(ℕ)]. ∀[g:K ⟶ J].
  (phi(g(f(rho))) phi(f ⋅ g(rho)) ∈ Point(face_lattice(K)))

Lemma: cubical-term-at-comp-is-1
[Gamma:j⊢]. ∀[phi:{Gamma ⊢ _:𝔽}]. ∀[I:fset(ℕ)]. ∀[rho:Gamma(I)]. ∀[J:fset(ℕ)]. ∀[f:J ⟶ I]. ∀[K:fset(ℕ)]. ∀[g:K ⟶ J].
  ((phi(f(rho)) 1 ∈ Point(face_lattice(J)))  (phi(f ⋅ g(rho)) 1 ∈ Point(face_lattice(K))))

Lemma: subset-context-1
[Gamma:j⊢]. sub_cubical_set{j:l}(Gamma, 1(𝔽); Gamma)

Lemma: context-1-subset
[Gamma:j⊢]. sub_cubical_set{j:l}(Gamma; Gamma, 1(𝔽))

Lemma: face-1-implies-subset
[Gamma:j⊢]. ∀[phi:{Gamma ⊢ _:𝔽}].  sub_cubical_set{j:l}(Gamma; Gamma, phi) supposing Gamma ⊢ (1(𝔽 phi)

Lemma: context-subset-1
[Gamma:j⊢]. {Gamma, 1(𝔽) ⊢ _} ≡ {Gamma ⊢ _}

Lemma: context-subset-1-subtype
[Gamma:j⊢]. ({Gamma, 1(𝔽) ⊢ _} ⊆{Gamma ⊢ _})

Lemma: context-subset-subtype-simple
[Gamma:j⊢]. ∀[phi:{Gamma ⊢ _:𝔽}].  ({Gamma ⊢ _} ⊆{Gamma, phi ⊢ _})

Lemma: thin-context-subset
[Gamma:j⊢]. ∀[phi:{Gamma ⊢ _:𝔽}]. ∀[t:{Gamma ⊢ _}].  Gamma, phi ⊢ t

Lemma: thin-context-subset-adjoin
[Gamma:j⊢]. ∀[phi:{Gamma ⊢ _:𝔽}]. ∀[T:{Gamma ⊢ _}]. ∀[t:{Gamma.T ⊢ _}].  Gamma, phi.T ⊢ t

Lemma: term-context-subset-subtype
[Gamma:j⊢]. ∀[phi1,phi2:{Gamma ⊢ _:𝔽}]. ∀[A:{Gamma ⊢ _}].
  {Gamma, phi1 ⊢ _:A} ⊆{Gamma, phi2 ⊢ _:A} supposing Gamma ⊢ (phi2  phi1)

Lemma: context-subset-term-subtype
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[phi:{Gamma ⊢ _:𝔽}].  ({Gamma ⊢ _:A} ⊆{Gamma, phi ⊢ _:A})

Lemma: respects-equality-context-subset-term
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[phi:{X ⊢ _:𝔽}].  respects-equality({X ⊢ _:A};{X, phi ⊢ _:A})

Lemma: face-term-implies-subtype
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[phi,psi:{Gamma ⊢ _:𝔽}].
  {Gamma, phi ⊢ _:A} ⊆{Gamma, psi ⊢ _:A} supposing Gamma ⊢ (psi  phi)

Lemma: context-iterated-subset
[Gamma:j⊢]. ∀[phi,psi:{Gamma ⊢ _:𝔽}].
  (sub_cubical_set{j:l}(Gamma, (phi ∧ psi); Gamma, psi, phi)
  ∧ sub_cubical_set{j:l}(Gamma, psi, phi; Gamma, (phi ∧ psi)))

Lemma: context-iterated-subset0
[X:j⊢]. ∀[xx,yy:{X ⊢ _:𝔽}].  sub_cubical_set{j:l}(X, xx, yy; X)

Lemma: context-iterated-subset1
[X:j⊢]. ∀[xx,yy:{X ⊢ _:𝔽}].  sub_cubical_set{j:l}(X, xx, yy; X, yy)

Lemma: context-iterated-subset2
[X:j⊢]. ∀[xx,yy:{X ⊢ _:𝔽}].  sub_cubical_set{j:l}(X, yy, xx; X, yy)

Lemma: context-subset-swap
[X:j⊢]. ∀[a,b:{X ⊢ _:𝔽}].  sub_cubical_set{j:l}(X, b, a; X, a, b)

Lemma: context-subset-term-1
[Gamma:j⊢]. ∀[T:{Gamma ⊢ _}].  ({Gamma ⊢ _:T} ⊆{Gamma, 1(𝔽) ⊢ _:T})

Lemma: context-subset-term-0
[Gamma:j⊢]. ∀[T:{Gamma ⊢ _}].  (Top ⊆{Gamma, 0(𝔽) ⊢ _:T})

Lemma: sub_cubical_set_functionality2
[Y,X:j⊢]. ∀[phi:{X ⊢ _:𝔽}].  sub_cubical_set{j:l}(Y, phi; X, phi) supposing sub_cubical_set{j:l}(Y; X)

Lemma: csm-ap-term-subset
[Gamma,K:j⊢]. ∀[s:K j⟶ Gamma]. ∀[phi:{Gamma ⊢ _:𝔽}]. ∀[A:{Gamma, phi ⊢ _}]. ∀[Z:{Gamma, phi ⊢ _:A}].
  ((Z)s ∈ {K, (phi)s ⊢ _:(A)s})

Lemma: csm-ap-term-subset-subset
[Gamma,K:j⊢]. ∀[s:K j⟶ Gamma]. ∀[phi:{Gamma ⊢ _:𝔽}]. ∀[psi:{Gamma, phi ⊢ _:𝔽}]. ∀[A:{Gamma, phi, psi ⊢ _}].
[Z:{Gamma, phi, psi ⊢ _:A}].
  ((Z)s ∈ {K, (phi)s, (psi)s ⊢ _:(A)s})

Lemma: context-subset-type-subtype
[G:j⊢]. ∀[phi:{G ⊢ _:𝔽}].  ({G ⊢_} ⊆{G, phi ⊢_})

Definition: constrained-cubical-term
{Gamma ⊢ _:A[phi |⟶ t]} ==  {a:{Gamma ⊢ _:A}| a ∈ {Gamma, phi ⊢ _:A}} 

Lemma: constrained-cubical-term_wf
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[phi:{Gamma ⊢ _:𝔽}]. ∀[t:{Gamma, phi ⊢ _:A}].  ({Gamma ⊢ _:A[phi |⟶ t]} ∈ 𝕌{[i j']})

Lemma: constrained-cubical-term-eqcd
[Gamma:j⊢]. ∀[A,B:{Gamma ⊢ _}]. ∀[phi,psi:{Gamma ⊢ _:𝔽}]. ∀[t:{Gamma, phi ⊢ _:A}]. ∀[t':{Gamma, psi ⊢ _:B}].
  ({Gamma ⊢ _:A[phi |⟶ t]} {Gamma ⊢ _:B[psi |⟶ t']} ∈ 𝕌{[i j']}) supposing 
     ((t' t ∈ {Gamma, phi ⊢ _:A}) and 
     (phi psi ∈ {Gamma ⊢ _:𝔽}) and 
     (A B ∈ {Gamma ⊢ _}))

Lemma: constrained-cubical-term-0
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[t:Top]. ∀[x:{Gamma ⊢ _:A}].  (x ∈ {Gamma ⊢ _:A[0(𝔽|⟶ t]})

Lemma: csm-context-subset-subtype
[Gamma:j⊢]. ∀[phi:{Gamma ⊢ _:𝔽}]. ∀[X:j⊢].  (X j⟶ Gamma, phi.𝕀 ⊆j⟶ Gamma.𝕀)

Lemma: subset-constrained-cubical-term
[X,Y:j⊢].
  ∀[A:{X ⊢ _}]. ∀[phi:{X ⊢ _:𝔽}]. ∀[t:{X, phi ⊢ _:A}].  ({X ⊢ _:A[phi |⟶ t]} ⊆{Y ⊢ _:A[phi |⟶ t]}) 
  supposing sub_cubical_set{j:l}(Y; X)

Lemma: csm-context-subset-subtype2
[Gamma:j⊢]. ∀[phi:{Gamma ⊢ _:𝔽}]. ∀[X:j⊢].  (Gamma j⟶ X ⊆Gamma, phi j⟶ X)

Lemma: csm-context-subset-subtype3
[Gamma:j⊢]. ∀[phi:{Gamma ⊢ _:𝔽}]. ∀[A:{Gamma ⊢ _}]. ∀[X:j⊢].  (Gamma.A ij⟶ X ⊆Gamma, phi.A ij⟶ X)

Lemma: csm-constrained-cubical-term
[Gamma,K:j⊢]. ∀[s:K j⟶ Gamma]. ∀[A:{Gamma ⊢ _}]. ∀[phi:{Gamma ⊢ _:𝔽}]. ∀[t:{Gamma, phi ⊢ _:A}].
[v:{Gamma ⊢ _:A[phi |⟶ t]}].
  ((v)s ∈ {K ⊢ _:(A)s[(phi)s |⟶ (t)s]})

Definition: same-cubical-type
Gamma ⊢ ==  B ∈ {Gamma ⊢ _}

Lemma: same-cubical-type_wf
[Gamma:j⊢]. ∀[A,B:{Gamma ⊢ _}].  (Gamma ⊢ B ∈ 𝕌{[i' j']})

Lemma: csm-same-cubical-type
[Gamma:j⊢]. ∀[A,B:{Gamma ⊢ _}].  ∀[H:j⊢]. ∀[s:H j⟶ Gamma].  H ⊢ (A)s (B)s supposing Gamma ⊢ B

Definition: same-cubical-term
X ⊢ u=v:A ==  v ∈ {X ⊢ _:A}

Lemma: same-cubical-term_wf
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[u,v:{X ⊢ _:A}].  (X ⊢ u=v:A ∈ 𝕌{[i j']})

Lemma: csm-same-cubical-term
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[u,v:{X ⊢ _:A}]. ∀[Y:j⊢]. ∀[s:Y j⟶ X].  Y ⊢ (u)s=(v)s:(A)s supposing X ⊢ u=v:A

Lemma: face-one-context-implies
[X:j⊢]. ∀[i:{X ⊢ _:𝕀}].  X, (i=1) ⊢ i=1(𝕀):𝕀

Lemma: face-zero-context-implies
[X:j⊢]. ∀[i:{X ⊢ _:𝕀}].  X, (i=0) ⊢ i=0(𝕀):𝕀

Lemma: same-cubical-type-0
[Gamma:j⊢]. ∀[A,B:{Gamma, 0(𝔽) ⊢ _}].  Gamma, 0(𝔽) ⊢ B

Lemma: same-cubical-type-zero-and-one
[G:j⊢]. ∀[A,B:{G, 0(𝔽) ⊢ _}]. ∀[i:{G ⊢ _:𝕀}].  G, ((i=0) ∧ (i=1)) ⊢ B

Lemma: subtype-context-subset-0
[X,Y:j⊢].  ({X ⊢ _} ⊆{Y, 0(𝔽) ⊢ _})

Lemma: same-cubical-type-by-cases
[Gamma:j⊢]. ∀[phi,psi:{Gamma ⊢ _:𝔽}]. ∀[A,B:{Gamma, (phi ∨ psi) ⊢ _}].
  (Gamma, (phi ∨ psi) ⊢ B) supposing (Gamma, phi ⊢ and Gamma, psi ⊢ B)

Lemma: same-cubical-term-by-cases
[Gamma:j⊢]. ∀[phi,psi:{Gamma ⊢ _:𝔽}]. ∀[A:{Gamma, (phi ∨ psi) ⊢ _}]. ∀[a,b:{Gamma, (phi ∨ psi) ⊢ _:A}].
  (Gamma, (phi ∨ psi) ⊢ a=b:A) supposing (Gamma, phi ⊢ a=b:A and Gamma, psi ⊢ a=b:A)

Lemma: same-cubical-term-by-cases2
[Gamma:j⊢]. ∀[phi,psi:{Gamma ⊢ _:𝔽}]. ∀[A:{Gamma ⊢ _}]. ∀[a,b:{Gamma, (phi ∨ psi) ⊢ _:A}].
  (Gamma, (phi ∨ psi) ⊢ a=b:A) supposing (Gamma, phi ⊢ a=b:A and Gamma, psi ⊢ a=b:A)

Lemma: same-cubical-type-by-list-cases
[Gamma:j⊢]. ∀L:{Gamma ⊢ _:𝔽List. ∀A,B:{Gamma, \/(L) ⊢ _}.  Gamma, \/(L) ⊢ supposing (∀phi∈L.Gamma, phi ⊢ B)

Lemma: context-subset-adjoin-subtype
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[phi:{Gamma ⊢ _:𝔽}].  ({Gamma.A ⊢ _} ⊆{Gamma, phi.A ⊢ _})

Lemma: face-forall-implies-0
[H:j⊢]. ∀[phi:{H.𝕀 ⊢ _:𝔽}]. ∀[X:Top].  H ⊢ ((∀ phi)  (phi)[0(𝕀)])

Lemma: face-forall-implies-1
[H:j⊢]. ∀[phi:{H.𝕀 ⊢ _:𝔽}]. ∀[X:Top].  H ⊢ ((∀ phi)  (phi)[1(𝕀)])

Lemma: face-forall-implies
[H:j⊢]. ∀[phi:{H.𝕀 ⊢ _:𝔽}].  H.𝕀 ⊢ (((∀ phi))p  phi)

Lemma: face-forall-implies-csm+
[H:j⊢]. ∀[phi:{H.𝕀 ⊢ _:𝔽}]. ∀[K:j⊢]. ∀[tau:K j⟶ H].  K.𝕀 ⊢ ((((∀ phi))tau)p  (phi)tau+)

Lemma: face-forall-decomp
[H:j⊢]. ∀[phi:{H.𝕀 ⊢ _:𝔽}].  H.𝕀 ⊢ (phi ⇐⇒ (((∀ phi))p ∨ ((phi ∧ (q=0)) ∨ (phi ∧ (q=1)))))

Lemma: face-forall-type-subtype
[H:j⊢]. ∀[phi:{H.𝕀 ⊢ _:𝔽}].  ({H.𝕀phi ⊢ _} ⊆{H.𝕀((∀ phi))p ⊢ _})

Lemma: csm-ap-term-wf-subset
[H:j⊢]. ∀[phi:{H ⊢ _:𝔽}]. ∀[A:{H ⊢ _}]. ∀[t:{H, phi ⊢ _:A}]. ∀[K:j⊢]. ∀[psi:{K ⊢ _:𝔽}]. ∀[B:{K ⊢ _}]. ∀[tau:K j⟶ H].
  ((t)tau ∈ {K, psi ⊢ _:B}) supposing (K, psi ⊢ (A)tau and K ⊢ (psi  (phi)tau))

Definition: partial-term-0
u[0] ==  (u)[0(𝕀)]

Lemma: partial-term-0_wf
[H:j⊢]. ∀[A:{H.𝕀 ⊢ _}]. ∀[phi:{H ⊢ _:𝔽}]. ∀[u:{H.𝕀(phi)p ⊢ _:A}].  (u[0] ∈ {H, phi ⊢ _:(A)[0(𝕀)]})

Definition: partial-term-1
u[1] ==  (u)[1(𝕀)]

Lemma: partial-term-1_wf
[H:j⊢]. ∀[A:{H.𝕀 ⊢ _}]. ∀[phi:{H ⊢ _:𝔽}]. ∀[u:{H.𝕀(phi)p ⊢ _:A}].  (u[1] ∈ {H, phi ⊢ _:(A)[1(𝕀)]})

Lemma: partial-term-1-subset
[H,xx,u:Top].  (partial-term-1(H, xx;u) partial-term-1(H;u))

Lemma: partial-term-0-subset
[H,xx,u:Top].  (partial-term-0(H, xx;u) partial-term-0(H;u))

Lemma: context-adjoin-subset0
[H:j⊢]. ∀[phi:{H ⊢ _:𝔽}].  ∀T:{H ⊢ _}. sub_cubical_set{k:l}(H.T, (phi)p; H, phi.T)

Lemma: context-adjoin-subset3
[H:j⊢]. ∀[phi:{H ⊢ _:𝔽}].  ∀T:{H ⊢ _}. sub_cubical_set{k:l}(H, phi.T; H.T, (phi)p)

Lemma: context-adjoin-subset4
[H:j⊢]. ∀[phi:{H ⊢ _:𝔽}].
  ∀T:{H ⊢_}. ∀[psi:{H.T ⊢ _:𝔽}]. ((psi (phi)p ∈ {H.T ⊢ _:𝔽})  sub_cubical_set{j:l}(H, phi.T; H.T, psi))

Lemma: context-adjoin-subset1
[H:j⊢]. ∀[phi:{H ⊢ _:𝔽}].  sub_cubical_set{j:l}(H.𝕀(phi)p; H, phi.𝕀)

Lemma: context-adjoin-subset2
[H:j⊢]. ∀[phi:{H ⊢ _:𝔽}].  sub_cubical_set{j:l}(H, phi.𝕀H.𝕀(phi)p)

Lemma: context-adjoin-subset
[X,Y:j⊢]. ∀[A:{Y ⊢ _}].  sub_cubical_set{[i j]:l}(X.A; Y.A) supposing sub_cubical_set{j:l}(X; Y)

Lemma: csm-subtype-iso-instance1
[X,H:j⊢]. ∀[phi:{H ⊢ _:𝔽}].  (X j⟶ H.𝕀(phi)p ⊆j⟶ H, phi.𝕀)

Lemma: csm-subtype-iso-instance2
[X,H:j⊢]. ∀[phi:{H ⊢ _:𝔽}].  (H.𝕀(phi)p j⟶ X ⊆H, phi.𝕀 j⟶ X)

Lemma: context-subset-map-equal
[H:j⊢]. ∀[phi:{H ⊢ _:𝔽}]. ∀[X:j⊢]. ∀[f,g:X j⟶ H.𝕀].  ((f g ∈ j⟶ H.𝕀 (f g ∈ X, ((phi)p)f j⟶ H, phi.𝕀))

Lemma: face-forall-subset
[Gamma,phi,xx:Top].  ((Gamma, xx ⊢ ∀ phi) (Gamma ⊢ ∀ phi))

Lemma: face-forall-map
[G,H:j⊢]. ∀[phi:{G ⊢ _:𝔽}]. ∀[s:H.𝕀 j⟶ G].  (s ∈ H, (∀ (phi)s).𝕀 j⟶ G, phi)

Lemma: csm+-ap-term-wf
[H,K:j⊢]. ∀[phi:{H ⊢ _:𝔽}]. ∀[A:{H, phi.𝕀 ⊢ _}]. ∀[tau:K j⟶ H]. ∀[t:{H, phi.𝕀 ⊢ _:A}].
  ((t)tau+ ∈ {K, (phi)tau.𝕀 ⊢ _:(A)tau+})

Lemma: csm-id-adjoin-subset
[Gamma,phi,u:Top].  ([u] [u])

Lemma: cubical-sigma-normalize
[X,A,B:Top].  (Σ ~ Σ B)

Lemma: cubical-sigma-subset
[X,A,B,phi:Top].  (Σ ~ Σ B)

Lemma: cubical-sigma-subset-adjoin
[X,A,B,T,phi:Top].  (Σ ~ Σ B)

Lemma: cubical-sigma-subset-adjoin2
[X,A,B,T1,T2,phi:Top].  (Σ ~ Σ B)

Lemma: term-p+0
[X,Y,Z,W,t,A:Top].  (((t)p+)[0(𝕀)] ((t)[0(𝕀)])p)

Lemma: term-p+1
[X,Y,Z,W,t,A:Top].  (((t)p+)[1(𝕀)] ((t)[1(𝕀)])p)

Lemma: face-1-in-context-subset
[G:j⊢]. ∀[phi:{G ⊢ _:𝔽}].  (phi 1(𝔽) ∈ {G, phi ⊢ _:𝔽})

Lemma: cubical-pi-context-subset
[X,phi,A,B:Top].  (X, phi ⊢ ΠX ⊢ ΠB)

Lemma: cubical-pi-subset
[X,phi,A,B,C:Top].  (X, phi.C ⊢ ΠX.C ⊢ ΠB)

Lemma: cubical-pi-subset-adjoin2
[X,phi,A,B,C,D:Top].  (X, phi.C.D ⊢ ΠX.C.D ⊢ ΠB)

Lemma: cubical-pi-subset-adjoin3
[X,phi,A,B,C,D,E:Top].  (X, phi.C.D.E ⊢ ΠX.C.D.E ⊢ ΠB)

Lemma: cubical-lambda-subset
[X,phi,b:Top].  ((λb) b))

Lemma: cubical-lambda-subset2
[X,A,phi,b:Top].  ((λb) b))

Lemma: implies-face-forall-holds
H:j⊢. ∀phi:{H.𝕀 ⊢ _:𝔽}.  (H.𝕀 ⊢ (1(𝔽 phi)  H ⊢ (1(𝔽 (∀ phi)))

Lemma: map-to-context-subset-disjoint
[Gamma:j⊢]. ∀[phi,psi:{Gamma ⊢ _:𝔽}]. ∀[H:j⊢].
  ∀[sigma:H j⟶ Gamma, (phi ∧ psi)]. ∀[I:fset(ℕ)].  H(I)) supposing Gamma ⊢ ((phi ∧ psi)  0(𝔽))

Comment: system type doc
================================= system type  =================================

System types allow definitions by cases.⋅

Definition: case-cube
case-cube(phi;A;B;I;rho) ==  if (phi(rho)==1) then A(rho) else B(rho) fi 

Lemma: case-cube_wf
[Gamma:j⊢]. ∀[phi,psi:{Gamma ⊢ _:𝔽}]. ∀[A:{Gamma, phi ⊢ _}]. ∀[B:{Gamma, psi ⊢ _}]. ∀[I:fset(ℕ)].
[rho:Gamma, (phi ∨ psi)(I)].
  (case-cube(phi;A;B;I;rho) ∈ Type)

Definition: case-type
(if phi then else B) ==  <λI,rho. case-cube(phi;A;B;I;rho), λI,J,f,rho,c. if (phi(rho)==1) then (c rho f) else (c rho \000Cf) fi >

Lemma: case-type_wf
[Gamma:j⊢]. ∀[phi,psi:{Gamma ⊢ _:𝔽}]. ∀[A:{Gamma, phi ⊢ _}]. ∀[B:{Gamma, psi ⊢ _}].
  Gamma, (phi ∨ psi) ⊢ (if phi then else B) supposing Gamma, (phi ∧ psi) ⊢ B

Lemma: case-type-partition
[Gamma:j⊢]. ∀[phi,psi:{Gamma ⊢ _:𝔽}].
  (∀[A:{Gamma, phi ⊢ _}]. ∀[B:{Gamma, psi ⊢ _}].  Gamma ⊢ (if phi then else B)) supposing 
     (Gamma ⊢ (1(𝔽 (phi ∨ psi)) and 
     Gamma ⊢ ((phi ∧ psi)  0(𝔽)))

Lemma: csm-case-type
[phi,A,B,s:Top].  (((if phi then else B))s (if (phi)s then (A)s else (B)s))

Lemma: case-type-same1
[Gamma:j⊢]. ∀[phi:{Gamma ⊢ _:𝔽}]. ∀[A:{Gamma, phi ⊢ _}]. ∀[psi:{Gamma ⊢ _:𝔽}]. ∀[B:{Gamma, psi ⊢ _}].
  Gamma, phi ⊢ (if phi then else B) A

Lemma: case-type-same2
[Gamma:j⊢]. ∀[phi:{Gamma ⊢ _:𝔽}]. ∀[A:{Gamma, phi ⊢ _}]. ∀[psi:{Gamma ⊢ _:𝔽}]. ∀[B:{Gamma, psi ⊢ _}].
  Gamma, psi ⊢ (if phi then else B) supposing Gamma, (phi ∧ psi) ⊢ B

Definition: compatible-system
compatible-system{i:l}(Gamma;sys) ==
  (∀phiA1,phiA2∈sys.  let phi1,A1 phiA1 
                      in let phi2,A2 phiA2 
                         in A1 A2 ∈ {Gamma, (phi1 ∧ phi2) ⊢ _})

Lemma: compatible-system_wf
[Gamma:j⊢]. ∀[sys:(phi:{Gamma ⊢ _:𝔽} × {Gamma, phi ⊢ _}) List].  (compatible-system{i:l}(Gamma;sys) ∈ ℙ{[i' j']})

Definition: system-type
system-type(sys) ==  reduce(λphi_A,B. let phi,A phi_A in (if phi then else B);discr(Top);sys)

Lemma: system-type-properties
Gamma:j⊢. ∀sys:(phi:{Gamma ⊢ _:𝔽} × {Gamma, phi ⊢ _}) List.
  (compatible-system{i:l}(Gamma;sys)
   (Gamma, \/(map(λp.(fst(p));sys)) ⊢ system-type(sys)
     ∧ (∀i:ℕ||sys||. (system-type(sys) (snd(sys[i])) ∈ {Gamma, fst(sys[i]) ⊢ _}))))

Lemma: system-type_wf
[Gamma:j⊢]. ∀[sys:(phi:{Gamma ⊢ _:𝔽} × {Gamma, phi ⊢ _}) List].
  Gamma, \/(map(λp.(fst(p));sys)) ⊢ system-type(sys) supposing compatible-system{i:l}(Gamma;sys)

Lemma: system-type-same
[Gamma:j⊢]. ∀[sys:(phi:{Gamma ⊢ _:𝔽} × {Gamma, phi ⊢ _}) List].
  ∀i:ℕ||sys||. Gamma, fst(sys[i]) ⊢ system-type(sys) snd(sys[i]) supposing compatible-system{i:l}(Gamma;sys)

Definition: case-term
(u ∨ v) ==  λI,rho. if (phi(rho)==1) then u(rho) else v(rho) fi 

Lemma: case-term_wf
[Gamma:j⊢]. ∀[phi,psi:{Gamma ⊢ _:𝔽}]. ∀[A:{Gamma, (phi ∨ psi) ⊢ _}]. ∀[u:{Gamma, phi ⊢ _:A}]. ∀[v:{Gamma, psi ⊢ _:A}].
  (u ∨ v) ∈ {Gamma, (phi ∨ psi) ⊢ _:A} supposing Gamma, (phi ∧ psi) ⊢ u=v:A

Lemma: case-term_wf2
[Gamma:j⊢]. ∀[phi,psi:{Gamma ⊢ _:𝔽}]. ∀[A:{Gamma, (phi ∨ psi) ⊢ _}]. ∀[v:{Gamma, psi ⊢ _:A}].
[u:{Gamma, phi ⊢ _:A[psi |⟶ v]}].
  ((u ∨ v) ∈ {Gamma, (phi ∨ psi) ⊢ _:A})

Lemma: case-term-same
[Gamma:j⊢]. ∀[phi:{Gamma ⊢ _:𝔽}]. ∀[A:{Gamma, phi ⊢ _}]. ∀[u:{Gamma, phi ⊢ _:A}].  ((u ∨ u) u ∈ {Gamma, phi ⊢ _:A})

Lemma: case-term-same2
[Gamma:j⊢]. ∀[phi,psi:{Gamma ⊢ _:𝔽}]. ∀[A:{Gamma ⊢ _}]. ∀[u:{Gamma, phi ⊢ _:A}]. ∀[v:{Gamma, psi ⊢ _:A}].
[w:{Gamma ⊢ _:A}].
  (Gamma, (phi ∨ psi) ⊢ (u ∨ v)=w:A) supposing (Gamma, phi ⊢ u=w:A and Gamma, psi ⊢ v=w:A)

Lemma: case-term-equal-left
[Gamma:j⊢]. ∀[phi:{Gamma ⊢ _:𝔽}]. ∀[A:{Gamma, phi ⊢ _}]. ∀[u:{Gamma, phi ⊢ _:A}]. ∀[v:Top].  Gamma, phi ⊢ (u ∨ v)=u:A

Lemma: case-term-equal-left'
[Gamma:j⊢]. ∀[phi:{Gamma ⊢ _:𝔽}]. ∀[A:{Gamma, phi ⊢ _}]. ∀[u:{Gamma, phi ⊢ _:A}]. ∀[v:Top]. ∀[x:{Gamma, phi ⊢ _:A}].
  Gamma, phi ⊢ (u ∨ v)=x:A supposing u ∈ {Gamma, phi ⊢ _:A}

Lemma: case-term-1
[Gamma:j⊢]. ∀[phi:{Gamma ⊢ _:𝔽}]. ∀[A:{Gamma ⊢ _}]. ∀[u:{Gamma ⊢ _:A}]. ∀[v:Top].
  Gamma ⊢ (u ∨ v)=u:A supposing phi 1(𝔽) ∈ {Gamma ⊢ _:𝔽}

Lemma: case-term-1'
[Gamma:j⊢]. ∀[phi:{Gamma ⊢ _:𝔽}]. ∀[A:{Gamma ⊢ _}]. ∀[u,x:{Gamma ⊢ _:A}]. ∀[v:Top].
  (Gamma ⊢ (u ∨ v)=x:A) supposing ((x u ∈ {Gamma ⊢ _:A}) and (phi 1(𝔽) ∈ {Gamma ⊢ _:𝔽}))

Lemma: case-term-equal-right
[Gamma:j⊢]. ∀[phi,psi:{Gamma ⊢ _:𝔽}]. ∀[A:{Gamma, (phi ∨ psi) ⊢ _}]. ∀[u:{Gamma, phi ⊢ _:A}]. ∀[v:{Gamma, psi ⊢ _:A}].
  Gamma, psi ⊢ (u ∨ v)=v:A supposing Gamma, (phi ∧ psi) ⊢ u=v:A

Lemma: case-term-equal-right'
[Gamma:j⊢]. ∀[phi,psi:{Gamma ⊢ _:𝔽}]. ∀[A:{Gamma, (phi ∨ psi) ⊢ _}]. ∀[u:{Gamma, phi ⊢ _:A}].
[v,x:{Gamma, psi ⊢ _:A}].
  (Gamma, psi ⊢ (u ∨ v)=x:A) supposing ((x v ∈ {Gamma, psi ⊢ _:A}) and Gamma, (phi ∧ psi) ⊢ u=v:A)

Lemma: case-term-equality-right
[Gamma:j⊢]. ∀[phi,psi:{Gamma ⊢ _:𝔽}]. ∀[A:{Gamma, (phi ∨ psi) ⊢ _}]. ∀[u:{Gamma, phi ⊢ _:A}]. ∀[v:{Gamma, psi ⊢ _:A}].
  ∀[x:{Gamma, (phi ∨ psi) ⊢ _:A}]. Gamma, psi ⊢ v=x:A supposing Gamma, (phi ∨ psi) ⊢ (u ∨ v)=x:A 
  supposing Gamma, (phi ∧ psi) ⊢ u=v:A

Lemma: case-term-wf3
[Gamma:j⊢]. ∀[phi,psi:{Gamma ⊢ _:𝔽}]. ∀[A:{Gamma, (phi ∨ psi) ⊢ _}]. ∀[v:{Gamma, psi ⊢ _:A}].
[u:{Gamma, phi ⊢ _:A[psi |⟶ v]}].
  ((u ∨ v) ∈ {t:{Gamma, (phi ∨ psi) ⊢ _:A}| Gamma, psi ⊢ t=v:A} )

Lemma: case-term-wf4
[Gamma:j⊢]. ∀[phi,psi:{Gamma ⊢ _:𝔽}]. ∀[A:{Gamma, phi ⊢ _}]. ∀[B:{Gamma, psi ⊢ _}]. ∀[u:{Gamma, phi ⊢ _:A}].
[v:{Gamma, psi ⊢ _:B}].
  ((u ∨ v) ∈ {Gamma, (phi ∨ psi) ⊢ _:(if phi then else B)}) supposing 
     (Gamma, (phi ∧ psi) ⊢ u=v:A and 
     Gamma, (phi ∧ psi) ⊢ B)

Lemma: csm-case-term
[phi,u,v,s:Top].  (((u ∨ v))s ((u)s ∨ (v)s))

Comment: metric space type doc
We define cubical type for any metric space that is ⌜WCE(X;d)⌝⋅

Definition: metric-type
metric-type(X;d) ==  <λI,rho. X, λI,J,f,rho,c. f>

Comment: cubical path type doc
================================= path type  ================================

The path type (Path_A b)  can now be defined as subspace of the free path
space Path(A) ==  (𝕀 ⟶ A). To do that we use general defintion of subset
of cubical type: {t:T | ∀I,alpha. psi[I;
                                        alpha;
                                        t]}.⋅

Comment: The Type (Path_A a b)
We used to define (Path_A b) using an equivalence class (quotient type)
of paths whose I-cubes were (I+z)-cubes of where was fresh coordinate
name (and such that z:=0 maps to and z:=1 maps to b).

But we later realized that we can start with the 
"free path space"  Path(A) == (𝕀 ⟶ A)
and define the "subspace" of paths such that p(0) and p(1) b.




Definition: cubical-type-restriction
cubical-type-restriction(X;T;I,a,x.psi[I;
                                       a;
                                       x]) ==
  ∀I,J:fset(ℕ). ∀f:J ⟶ I. ∀a:X(I). ∀x:T(a).
    (psi[I;
         a;
         x]
     psi[J;
           f(a);
           (x f)])

Lemma: cubical-type-restriction_wf
[X:j⊢]. ∀[T:{X ⊢ _}]. ∀[psi:I:fset(ℕ) ⟶ alpha:X(I) ⟶ T(alpha) ⟶ ℙ{[i' j']}].
  (cubical-type-restriction(X;T;I,a,t.psi[I;a;t]) ∈ ℙ{[i' j']})

Lemma: cubical-type-restriction-and
[X:j⊢]. ∀[T:{X ⊢ _}]. ∀[psi,phi:I:fset(ℕ) ⟶ alpha:X(I) ⟶ T(alpha) ⟶ ℙ].
  (cubical-type-restriction(X;T;I,a,t.psi[I;a;t])
   cubical-type-restriction(X;T;I,a,t.phi[I;a;t])
   cubical-type-restriction(X;T;I,a,t.psi[I;a;t] ∧ phi[I;a;t]))

Lemma: cubical-type-restriction-eq
[X:j⊢]. ∀[T,A:{X ⊢ _}]. ∀[a:{X ⊢ _:A}]. ∀[g:I:fset(ℕ) ⟶ alpha:X(I) ⟶ T(alpha) ⟶ A(alpha)].
  cubical-type-restriction(X;T;I,alpha,t.(g alpha t) a(alpha) ∈ A(alpha)) 
  supposing ∀I,J:fset(ℕ). ∀f:J ⟶ I. ∀alpha:X(I). ∀t:T(alpha).
              ((g f(alpha) (t alpha f)) (g alpha alpha f) ∈ A(f(alpha)))

Definition: cubical-subset
{t:T | ∀I,alpha. psi[I;
                     alpha;
                     t]} ==
  <λI,alpha. {t:T(alpha)| 
              psi[I;
                  alpha;
                  t]} 
  , λI,J,f,alpha,p. (p alpha f)
  >

Lemma: cubical-subset_wf
[X:j⊢]. ∀[T:{X ⊢ _}]. ∀[psi:I:fset(ℕ) ⟶ alpha:X(I) ⟶ T(alpha) ⟶ ℙ].
  X ⊢ {t:T | ∀I,alpha. psi[I;alpha;t]} supposing cubical-type-restriction(X;T;I,a,t.psi[I;a;t])

Lemma: cubical-subset-term
[X:j⊢]. ∀[T:{X ⊢ _}]. ∀[psi:I:fset(ℕ) ⟶ alpha:X(I) ⟶ T(alpha) ⟶ ℙ].
  ∀[t:{X ⊢ _:T}]
    t ∈ {X ⊢ _:{t:T | ∀I,alpha. psi[I;alpha;t]}} supposing ∀I:fset(ℕ). ∀alpha:X(I).  psi[I;alpha;t(alpha)] 
  supposing cubical-type-restriction(X;T;I,a,t.psi[I;a;t])

Lemma: csm-cubical-subset
[T,psi,s:Top].  (({t:T | ∀I,alpha. psi[I;alpha;t]})s {t:(T)s | ∀I,alpha. psi[I;(s)alpha;t]})

Definition: pathtype
Path(A) ==  (𝕀 ⟶ A)

Lemma: pathtype_wf
[X:j⊢]. ∀[A:{X ⊢ _}].  X ⊢ Path(A)

Definition: cubicalpath-app
pth ==  app(pth; r)

Lemma: cubicalpath-app_wf
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[t:{X ⊢ _:Path(A)}]. ∀[r:{X ⊢ _:𝕀}].  (t r ∈ {X ⊢ _:A})

Lemma: path-restriction
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[a,b:{X ⊢ _:A}].
  cubical-type-restriction(X;Path(A);I,a1,p.((p 0) a(a1) ∈ A(a1)) ∧ ((p 1) b(a1) ∈ A(a1)))

Definition: path-type
(Path_A b) ==  {p:Path(A) | ∀I,alpha. ((p 0) a(alpha) ∈ A(alpha)) ∧ ((p 1) b(alpha) ∈ A(alpha))}

Lemma: path-type_wf
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[a,b:{X ⊢ _:A}].  X ⊢ (Path_A b)

Lemma: path-type-at-subtype
X:j⊢. ∀A:{X ⊢ _}. ∀a,b:{X ⊢ _:A}. ∀I:fset(ℕ). ∀rho:X(I).  ((Path_A b)(rho) ⊆Path(A)(rho))

Lemma: path-type-subtype
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[a,b:{X ⊢ _:A}].  ({X ⊢ _:(Path_A b)} ⊆{X ⊢ _:Path(A)})

Lemma: pathtype-subset
[X,A,phi:Top].  (Path(A) Path(A))

Lemma: path-type-subset
[X,A,a,b,phi:Top].  ((X, phi ⊢ Path_A b) (X ⊢ Path_A b))

Lemma: path-type-subset-adjoin
[X,T,A,w,a,phi:Top].  ((X, phi.T ⊢ Path_A w) (X.T ⊢ Path_A w))

Lemma: path-type-subset-adjoin2
[X,T1,T2,A,w,a,phi:Top].  ((X, phi.T1.T2 ⊢ Path_A w) (X.T1.T2 ⊢ Path_A w))

Lemma: path-type-subset-adjoin3
[X,T1,T2,T3,A,w,a,phi:Top].  ((X, phi.T1.T2.T3 ⊢ Path_A w) (X.T1.T2.T3 ⊢ Path_A w))

Lemma: path-type-subset-adjoin4
[X,T1,T2,T3,T4,A,w,a,phi:Top].  ((X, phi.T1.T2.T3.T4 ⊢ Path_A w) (X.T1.T2.T3.T4 ⊢ Path_A w))

Lemma: path-type-ap-morph
[X,A,a,b,I,J,f,alpha,u:Top].  ((u alpha f) ~ λK,g. (u f ⋅ g))

Lemma: path-type-at
[X,A,a,b,I,alpha:Top].
  ((Path_A b)(alpha) {p:{w:J:fset(ℕ) ⟶ f:J ⟶ I ⟶ u:Point(dM(J)) ⟶ A(f(alpha))| 
                             ∀J,K:fset(ℕ). ∀f:J ⟶ I. ∀g:K ⟶ J. ∀u:Point(dM(J)).
                               ((w f(alpha) g) (w f ⋅ (u f(alpha) g)) ∈ A(g(f(alpha))))} 
                          ((p 0) a(alpha) ∈ A(alpha)) ∧ ((p 1) b(alpha) ∈ A(alpha))} )

Lemma: csm-pathtype
X,Delta:j⊢. ∀s:Delta j⟶ X. ∀A:{X ⊢ _}.  ((Path(A))s Path((A)s) ∈ {Delta ⊢ _})

Lemma: csm-path-type
X,Delta:j⊢. ∀s:Delta j⟶ X. ∀A:{X ⊢ _}. ∀a,b:{X ⊢ _:A}.
  (((Path_A b))s (Delta ⊢ Path_(A)s (a)s (b)s) ∈ {Delta ⊢ _})

Lemma: csm-path-type-id-adjoin
[X:j⊢]. ∀[A,B:{X ⊢ _}]. ∀[a,b:{X.B ⊢ _:(A)p}]. ∀[u:{X ⊢ _:B}].
  (((Path_(A)p b))[u] (X ⊢ Path_A (a)[u] (b)[u]) ∈ {X ⊢ _})

Lemma: path-type-q-id-adjoin
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[a,u:{X ⊢ _:A}].  (((Path_(A)p (a)p q))[u] (X ⊢ Path_A u) ∈ {X ⊢ _})

Lemma: path-type-q-csm-adjoin
[X,H:j⊢]. ∀[s:H j⟶ X]. ∀[A:{X ⊢ _}]. ∀[a:{X ⊢ _:A}]. ∀[u:{H ⊢ _:(A)s}].
  (((Path_(A)p (a)p q))(s;u) (H ⊢ Path_(A)s (a)s u) ∈ {H ⊢ _})

Lemma: path-type-p
[X:j⊢]. ∀[A,B:{X ⊢ _}]. ∀[a,b:{X ⊢ _:A}].  ((X.B ⊢ Path_(A)p (a)p (b)p) ((Path_A b))p ∈ {X.B ⊢ _})

Definition: path-eta
path-eta(pth) ==  (pth)p q

Lemma: path-eta_wf
[G:j⊢]. ∀[A:{G ⊢ _}]. ∀[pth:{G ⊢ _:Path(A)}].  (path-eta(pth) ∈ {G.𝕀 ⊢ _:(A)p})

Definition: path-elim
path-elim(pth) ==  (pth)p q

Lemma: path-elim_wf
[G:j⊢]. ∀[A:{G ⊢ _}]. ∀[pth:{G ⊢ _:Path(A)}].  (path-elim(pth) ∈ {G.𝕀 ⊢ _:(A)p})

Definition: cubical-path-app
pth ==  pth r

Lemma: cubical-path-app_wf
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[a,b:{X ⊢ _:A}]. ∀[t:{X ⊢ _:(Path_A b)}]. ∀[r:{X ⊢ _:𝕀}].  (t r ∈ {X ⊢ _:A})

Lemma: cubical-path-app-sq
[t,r:Top].  (t r)

Lemma: csm-cubicalpath-app
[t,r,s:Top].  ((t r)s (t)s (r)s)

Lemma: path-eta-id-adjoin
[G,u,pth:Top].  ((path-eta(pth))[u] pth u)

Lemma: path-eta-0
[G,pth:Top].  ((path-eta(pth))[0(𝕀)] pth 0(𝕀))

Lemma: path-eta-1
[G,pth:Top].  ((path-eta(pth))[1(𝕀)] pth 1(𝕀))

Lemma: csm-cubical-path-app
[t,r,s:Top].  ((t r)s (t)s (r)s)

Lemma: csm-path-eta
[pth,s:Top].  ((path-eta(pth))s ((pth)p)s (q)s)

Lemma: cubical-path-app-0
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[a,b:{X ⊢ _:A}]. ∀[t:{X ⊢ _:(Path_A b)}].  (t 0(𝕀a ∈ {X ⊢ _:A})

Lemma: cubical-path-app-1
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[a,b:{X ⊢ _:A}]. ∀[t:{X ⊢ _:(Path_A b)}].  (t 1(𝕀b ∈ {X ⊢ _:A})

Lemma: cubical-path-ap-id-adjoin
[pth,x,Z:Top].  (((pth)p q)[x] pth x)

Lemma: cubical-path-ap-id-adjoin2
[pth,x,Z:Top].  (((pth)p q)[x] pth x)

Lemma: path-type-ext-eq
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[a,b:{X ⊢ _:A}].
  {t:{X ⊢ _:Path(A)}| (t 0(𝕀a ∈ {X ⊢ _:A}) ∧ (t 1(𝕀b ∈ {X ⊢ _:A})}  ≡ {X ⊢ _:(Path_A b)}

Lemma: path-type-sub-pathtype
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[a,b:{X ⊢ _:A}].  ({X ⊢ _:(Path_A b)} ⊆{X ⊢ _:Path(A)})

Lemma: csm-path-type-sub-pathtype
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[a,b:{X ⊢ _:A}]. ∀[H:j⊢]. ∀[tau:H j⟶ X].  ({H ⊢ _:((Path_A b))tau} ⊆{H ⊢ _:(Path(A))tau})

Lemma: paths-equal
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[a,b:{X ⊢ _:A}]. ∀[p:{X ⊢ _:(Path_A b)}]. ∀[q:{X ⊢ _:Path(A)}].
  q ∈ {X ⊢ _:(Path_A b)} supposing q ∈ {X ⊢ _:Path(A)}

Lemma: csm-paths-equal
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[a,b:{X ⊢ _:A}]. ∀[p:{X ⊢ _:(Path_A b)}]. ∀[H:j⊢]. ∀[tau:H j⟶ X]. ∀[q:{H ⊢ _:(Path(A))tau}].
  (p)tau q ∈ {H ⊢ _:((Path_A b))tau} supposing (p)tau q ∈ {H ⊢ _:(Path(A))tau}

Definition: term-to-path
<>(a) ==  a)

Lemma: term-to-path_wf
[X:j⊢]. ∀[A:{X ⊢ _}].  ∀a:{X.𝕀 ⊢ _:(A)p}. (<>(a) ∈ {X ⊢ _:(Path_A (a)[0(𝕀)] (a)[1(𝕀)])})

Lemma: term-to-path-subset
[X,phi,a:Top].  (X, phi ⊢ <>(a) X ⊢ <>(a))

Lemma: csm-term-to-path
[G:j⊢]. ∀[A:{G ⊢ _}].
  ∀a:{G.𝕀 ⊢ _:(A)p}
    ∀[H:j⊢]. ∀[sigma:H j⟶ G].
      ((<>(a))sigma H ⊢ <>((a)sigma+) ∈ {H ⊢ _:(Path_(A)sigma ((a)sigma+)[0(𝕀)] ((a)sigma+)[1(𝕀)])})

Lemma: term-to-path-beta
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[b:{X.𝕀 ⊢ _:(A)p}]. ∀[r:{X ⊢ _:𝕀}].  (<>(b) (b)[r] ∈ {X ⊢ _:A})

Lemma: term-to-path-eta
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[a,b:{X ⊢ _:A}]. ∀[pth:{X ⊢ _:(Path_A b)}].  (X ⊢ <>((pth)p q) pth ∈ {X ⊢ _:(Path_A b)})

Lemma: equal-paths-eta
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[p,q:{X ⊢ _:Path(A)}].
  q ∈ {X ⊢ _:Path(A)} supposing path-eta(p) path-eta(q) ∈ {X.𝕀 ⊢ _:(A)p}

Lemma: paths-equal-eta
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[a,b:{X ⊢ _:A}]. ∀[p:{X ⊢ _:(Path_A b)}]. ∀[q:{X ⊢ _:Path(A)}].
  q ∈ {X ⊢ _:(Path_A b)} supposing path-eta(p) path-eta(q) ∈ {X.𝕀 ⊢ _:(A)p}

Lemma: term-to-path-app-snd
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[a:{X ⊢ _:A}].  ((<>((a)p))p (a)p ∈ {X.𝕀 ⊢ _:(A)p})

Lemma: term-to-path-wf
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[u,v:{X ⊢ _:A}].
  ∀a:{X.𝕀 ⊢ _:(A)p}. (<>(a) ∈ {X ⊢ _:(Path_A v)}) supposing (X ⊢ (a)[0(𝕀)]=u:A and X ⊢ (a)[1(𝕀)]=v:A)

Lemma: term-to-path-equal
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[u,v:{X ⊢ _:A}].
  ∀a:{X.𝕀 ⊢ _:(A)p}
    (∀[b:{X.𝕀 ⊢ _:(A)p}]. X ⊢ <>(a) X ⊢ <>(b) ∈ {X ⊢ _:(Path_A v)} supposing b ∈ {X.𝕀 ⊢ _:(A)p}) supposing 
       (X ⊢ (a)[0(𝕀)]=u:A and 
       X ⊢ (a)[1(𝕀)]=v:A)

Definition: term-to-pathtype
<>==  <>(a)

Lemma: term-to-pathtype_wf
[X:j⊢]. ∀[A:{X ⊢ _}].  ∀a:{X.𝕀 ⊢ _:(A)p}. (<>a ∈ {X ⊢ _:Path(A)})

Lemma: term-to-pathtype-eta
[X:j⊢]. ∀[A:{X ⊢ _}].  ∀pth:{X ⊢ _:Path(A)}. (<>(pth)p pth ∈ {X ⊢ _:Path(A)})

Lemma: csm-term-to-pathtype
[G:j⊢]. ∀[A:{G ⊢ _}].
  ∀a:{G.𝕀 ⊢ _:(A)p}. ∀[H:j⊢]. ∀[sigma:H j⟶ G].  ((<>a)sigma = <>(a)sigma+ ∈ {H ⊢ _:Path((A)sigma)})

Lemma: term-to-pathtype-beta
[G:j⊢]. ∀[A:{G ⊢ _}].  ∀r:{G ⊢ _:𝕀}. ∀a:{G.𝕀 ⊢ _:(A)p}.  (<>(a)[r] ∈ {G ⊢ _:A})

Definition: map-path
map-path(G;f;pth) ==  <>(app((f)p; path-eta(pth)))

Lemma: map-path_wf
[G:j⊢]. ∀[S,T:{G ⊢ _}]. ∀[f:{G ⊢ _:(S ⟶ T)}]. ∀[a,b:{G ⊢ _:S}]. ∀[pth:{G ⊢ _:(Path_S b)}].
  (map-path(G;f;pth) ∈ {G ⊢ _:(Path_T app(f; a) app(f; b))})

Definition: cubical-refl
refl(a) ==  <>((a)p)

Lemma: cubical-refl_wf
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[a:{X ⊢ _:A}].  (refl(a) ∈ {X ⊢ _:(Path_A a)})

Lemma: term-to-path-is-refl
[X:j⊢]. ∀[A:{X ⊢ _}].  ∀a:{X ⊢ _:A}. (X ⊢ <>((a)p) refl(a) ∈ {X ⊢ _:(Path_A a)})

Lemma: csm-cubical-refl
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[a:{X ⊢ _:A}]. ∀[H:j⊢]. ∀[s:H j⟶ X].  ((refl(a))s refl((a)s) ∈ {H ⊢ _:(Path_(A)s (a)s (a)s)})

Lemma: cubical-refl-p
[X:j⊢]. ∀[A,B:{X ⊢ _}]. ∀[a:{X ⊢ _:A}].  (refl((a)p) (refl(a))p ∈ {X.B ⊢ _:((Path_A a))p})

Lemma: cubical-refl-p-p
[X:j⊢]. ∀[A,B:{X ⊢ _}]. ∀[a:{X ⊢ _:A}]. ∀[C:{X.B ⊢ _}].
  (refl(((a)p)p) ((refl(a))p)p ∈ {X.B.C ⊢ _:(((Path_A a))p)p})

Lemma: refl-path-app
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[a:{X ⊢ _:A}]. ∀[r:{X ⊢ _:𝕀}].  (refl(a) a ∈ {X ⊢ _:A})

Lemma: cubical-refl-app-snd
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[a:{X ⊢ _:A}].  ((refl(a))p (a)p ∈ {X.𝕀 ⊢ _:(A)p})

Definition: rev-path
rev-path(G;pth) ==  <>((pth)p 1-(q))

Lemma: rev-path_wf
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[b,a:{X ⊢ _:A}]. ∀[pth:{X ⊢ _:(Path_A b)}].  (rev-path(X;pth) ∈ {X ⊢ _:(Path_A a)})

Definition: contractible-type
Contractible(A) ==  Σ A Π(A)p (Path_((A)p)p (q)p q)

Lemma: contractible-type_wf
[X:j⊢]. ∀[A:{X ⊢ _}].  X ⊢ Contractible(A)

Lemma: contractible-type-subset
[X,A,phi:Top].  (X, phi ⊢ Contractible(A) X ⊢ Contractible(A))

Lemma: contractible-type-subset2
[X,A,B,C,phi:Top].  (X, phi.B.C ⊢ Contractible(A) X.B.C ⊢ Contractible(A))

Lemma: csm-contractible-type
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[Z:j⊢]. ∀[s:Z j⟶ X].  ((Contractible(A))s Z ⊢ Contractible((A)s) ∈ {Z ⊢ _})

Lemma: contractible-type-at
[X,A,I,rho:Top].  (Contractible(A)(rho) u:A(rho) × cubical-pi-family(X.A;(A)p;(Path_((A)p)p (q)p q);I;(rho;u)))

Definition: contr-center
contr-center(c) ==  c.1

Lemma: contr-center_wf
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[c:{X ⊢ _:Contractible(A)}].  (contr-center(c) ∈ {X ⊢ _:A})

Lemma: csm-contr-center
[c,s:Top].  ((contr-center(c))s contr-center((c)s))

Definition: contr-path
contr-path(c;x) ==  app(c.2; x)

Lemma: contr-path_wf
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[c:{X ⊢ _:Contractible(A)}]. ∀[x:{X ⊢ _:A}].
  (contr-path(c;x) ∈ {X ⊢ _:(Path_A contr-center(c) x)})

Lemma: csm-contr-path
[s,c,x:Top].  ((contr-path(c;x))s contr-path((c)s;(x)s))

Definition: contr-witness
contr-witness(X;c;p) ==  cubical-pair(c;(λp))

Lemma: contr-witness_wf
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[c:{X ⊢ _:A}]. ∀[p:{X.A ⊢ _:(Path_(A)p (c)p q)}].
  (contr-witness(X;c;p) ∈ {X ⊢ _:Contractible(A)})

Definition: cubical-fiber
Fiber(w;a) ==  Σ (Path_(A)p (a)p app((w)p; q))

Lemma: cubical-fiber_wf
[X:j⊢]. ∀[T,A:{X ⊢ _}]. ∀[w:{X ⊢ _:(T ⟶ A)}]. ∀[a:{X ⊢ _:A}].  X ⊢ Fiber(w;a)

Lemma: cubical-fiber-subset
[X,T,A,w,a,phi:Top].  (X, phi ⊢ Fiber(w;a) X ⊢ Fiber(w;a))

Lemma: cubical-fiber-subset-adjoin
[X,T,A,B,w,a,phi:Top].  (X, phi.B ⊢ Fiber(w;a) X.B ⊢ Fiber(w;a))

Lemma: cubical-fiber-subset-adjoin2
[X,T,A,B1,B2,w,a,phi:Top].  (X, phi.B1.B2 ⊢ Fiber(w;a) X.B1.B2 ⊢ Fiber(w;a))

Lemma: csm-cubical-fiber
[X:j⊢]. ∀[T,A:{X ⊢ _}]. ∀[w:{X ⊢ _:(T ⟶ A)}]. ∀[a:{X ⊢ _:A}]. ∀[Z:j⊢]. ∀[s:Z j⟶ X].
  ((Fiber(w;a))s Z ⊢ Fiber((w)s;(a)s) ∈ {Z ⊢ _})

Lemma: cubical-fiber-at
[X,T,A,w,a,I,rho:Top].  (Fiber(w;a)(rho) u:T(rho) × (Path_(A)p (a)p app((w)p; q))((rho;u)))

Definition: fiber-point
fiber-point(t;c) ==  cubical-pair(t;c)

Lemma: fiber-point_wf
[X:j⊢]. ∀[T,A:{X ⊢ _}]. ∀[f:{X ⊢ _:(T ⟶ A)}]. ∀[a:{X ⊢ _:A}]. ∀[t:{X ⊢ _:T}]. ∀[c:{X ⊢ _:(Path_A app(f; t))}].
  (fiber-point(t;c) ∈ {X ⊢ _:Fiber(f;a)})

Lemma: csm-fiber-point
[s,t,c:Top].  ((fiber-point(t;c))s fiber-point((t)s;(c)s))

Lemma: fiber-subset
[X,T,A,f,a,phi:Top].  (X, phi ⊢ Fiber(f;a) X ⊢ Fiber(f;a))

Definition: fiber-member
fiber-member(p) ==  p.1

Lemma: fiber-member_wf
[X:j⊢]. ∀[T,A:{X ⊢ _}]. ∀[w:{X ⊢ _:(T ⟶ A)}]. ∀[a:{X ⊢ _:A}]. ∀[p:{X ⊢ _:Fiber(w;a)}].  (fiber-member(p) ∈ {X ⊢ _:T})

Lemma: csm-fiber-member
[X:j⊢]. ∀[T,A:{X ⊢ _}]. ∀[w:{X ⊢ _:(T ⟶ A)}]. ∀[a:{X ⊢ _:A}]. ∀[p:{X ⊢ _:Fiber(w;a)}]. ∀[H:j⊢]. ∀[s:H j⟶ X].
  ((fiber-member(p))s fiber-member((p)s) ∈ {H ⊢ _:(T)s})

Lemma: fiber-member-fiber-point
[X,a,b:Top].  (fiber-member(fiber-point(a;b)) (a)1(X))

Definition: fiber-path
fiber-path(p) ==  p.2

Lemma: fiber-path_wf
[X:j⊢]. ∀[T,A:{X ⊢ _}]. ∀[w:{X ⊢ _:(T ⟶ A)}]. ∀[a:{X ⊢ _:A}]. ∀[p:{X ⊢ _:Fiber(w;a)}].
  (fiber-path(p) ∈ {X ⊢ _:(Path_A app(w; fiber-member(p)))})

Lemma: csm-fiber-path
[X:j⊢]. ∀[T,A:{X ⊢ _}]. ∀[w:{X ⊢ _:(T ⟶ A)}]. ∀[a:{X ⊢ _:A}]. ∀[p:{X ⊢ _:Fiber(w;a)}]. ∀[H:j⊢]. ∀[s:H j⟶ X].
  ((fiber-path(p))s fiber-path((p)s) ∈ {H ⊢ _:(Path_(A)s (a)s app((w)s; fiber-member((p)s)))})

Lemma: fiber-path-fiber-point
[X,a,b:Top].  (fiber-path(fiber-point(a;b)) (b)1(X))

Lemma: sigma-path-fst
X:j⊢. ∀A:{X ⊢ _}. ∀B:{X.A ⊢ _}. ∀x,y:{X ⊢ _:Σ B}.  ({X ⊢ _:(Path_Σ y)}  {X ⊢ _:(Path_A x.1 y.1)})

Definition: is-cubical-equiv
IsEquiv(T;A;w) ==  ΠContractible(Fiber((w)p;q))

Lemma: is-cubical-equiv_wf
[X:j⊢]. ∀[T,A:{X ⊢ _}]. ∀[w:{X ⊢ _:(T ⟶ A)}].  X ⊢ IsEquiv(T;A;w)

Lemma: csm-is-cubical-equiv
[X:j⊢]. ∀[T,A:{X ⊢ _}]. ∀[w:{X ⊢ _:(T ⟶ A)}]. ∀[Z:j⊢]. ∀[s:Z j⟶ X].
  ((IsEquiv(T;A;w))s IsEquiv((T)s;(A)s;(w)s) ∈ {Z ⊢ _})

Lemma: is-cubical-equiv-at
[X:j⊢]. ∀[T:{X ⊢ _}]. ∀[A,w,I,rho:Top].
  (IsEquiv(T;A;w)(rho) {eqv:J:fset(ℕ)
                          ⟶ f:J ⟶ I
                          ⟶ u:A(f(rho))
                          ⟶ (u1:u1:T((p)(f(rho);u)) × (Path_((A)p)p (q)p app(((w)p)p; q))(((f(rho);u);u1))
                             × cubical-pi-family(X.A.Fiber((w)p;q);(Fiber((w)p;q))p;(Path_((Fiber((w)p;q))p)p (q)p
                                                                                          q);J;((f(rho);u);u1)))| 
                          ∀J,K:fset(ℕ). ∀f:J ⟶ I. ∀g:K ⟶ J. ∀u:A(f(rho)).
                            ((eqv (f(rho);u) g)
                            (eqv f ⋅ (u f(rho) g))
                            ∈ (u1:u1:T((p)g((f(rho);u))) × (Path_((A)p)p (q)p app(((w)p)p; q))((g((f(rho);u));u1))
                              × cubical-pi-family(X.A.Fiber((w)p;q);(Fiber((w)p;q))p;(Path_((Fiber((w)p;q))p)p (q)p
                                                                                           q);K;(g((f(rho);u));u1))))} )

Definition: is-equiv-witness
is-equiv-witness(G;A;c;p) ==  contr-witness(G.A;c;p))

Lemma: is-equiv-witness_wf
[X:j⊢]. ∀[T,A:{X ⊢ _}]. ∀[w:{X ⊢ _:(T ⟶ A)}]. ∀[c:{X.A ⊢ _:Fiber((w)p;q)}].
[p:{X.A.Fiber((w)p;q) ⊢ _:(Path_(Fiber((w)p;q))p (c)p q)}].
  (is-equiv-witness(X;A;c;p) ∈ {X ⊢ _:IsEquiv(T;A;w)})

Definition: cubical-equiv
Equiv(T;A) ==  Σ (T ⟶ A) IsEquiv((T)p;(A)p;q)

Lemma: cubical-equiv_wf
[X:j⊢]. ∀[T,A:{X ⊢ _}].  X ⊢ Equiv(T;A)

Lemma: csm-cubical-equiv
[H,K:j⊢]. ∀[tau:K j⟶ H]. ∀[A,E:{H ⊢ _}].  ((Equiv(A;E))tau Equiv((A)tau;(E)tau) ∈ {K ⊢ _})

Lemma: cubical-equiv-p
[H:j⊢]. ∀[T,A,E:{H ⊢ _}].  ((Equiv(A;E))p Equiv((A)p;(E)p) ∈ {H.T ⊢ _})

Lemma: equal-cubical-equiv-at
[X:j⊢]. ∀[T,A:{X ⊢ _}]. ∀[I:fset(ℕ)]. ∀[a:X(I)]. ∀[x:Equiv(T;A)(a)]. ∀[y:u:(T ⟶ A)(a) × IsEquiv((T)p;(A)p;q)((a;u))].
  y ∈ Equiv(T;A)(a) supposing y ∈ (u:(T ⟶ A)(a) × IsEquiv((T)p;(A)p;q)((a;u)))

Lemma: cubical-equiv-subset
[X,T,A,phi:Top].  (Equiv(T;A) Equiv(T;A))

Definition: equiv-fun
equiv-fun(f) ==  f.1

Lemma: equiv-fun_wf
[G:j⊢]. ∀[A,T:{G ⊢ _}]. ∀[f:{G ⊢ _:Equiv(T;A)}].  (equiv-fun(f) ∈ {G ⊢ _:(T ⟶ A)})

Lemma: csm-equiv-fun
[s,f:Top].  ((equiv-fun(f))s equiv-fun((f)s))

Definition: equiv-contr
equiv-contr(f;a) ==  app(f.2; a)

Lemma: equiv-contr_wf
[G:j⊢]. ∀[A,T:{G ⊢ _}]. ∀[f:{G ⊢ _:Equiv(T;A)}]. ∀[a:{G ⊢ _:A}].
  (equiv-contr(f;a) ∈ {G ⊢ _:Contractible(Fiber(equiv-fun(f);a))})

Lemma: csm-equiv-contr
[s,f,a:Top].  ((equiv-contr(f;a))s equiv-contr((f)s;(a)s))

Definition: equiv-witness
equiv-witness(f;cntr) ==  cubical-pair(f;cntr)

Lemma: equiv-witness_wf
[G:j⊢]. ∀[A,T:{G ⊢ _}]. ∀[f:{G ⊢ _:(T ⟶ A)}]. ∀[iseq:{G ⊢ _:IsEquiv(T;A;f)}].
  (equiv-witness(f;iseq) ∈ {G ⊢ _:Equiv(T;A)})

Lemma: cubical-fiber-id-fun
X:j⊢. ∀T:{X ⊢ _}.  ∀[u:{X ⊢ _:T}]. (X ⊢ Fiber(cubical-id-fun(X);u) = Σ (Path_(T)p (u)p q) ∈ {X ⊢ _})

Definition: path-point
path-point(pth) ==  (pth)p q

Lemma: path-point_wf
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[a,b:{X ⊢ _:A}]. ∀[pth:{X ⊢ _:(Path_A b)}].  (path-point(pth) ∈ {X.𝕀 ⊢ _:(A)p})

Lemma: path-point-0
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[a,b:{X ⊢ _:A}]. ∀[pth:{X ⊢ _:(Path_A b)}].  ((path-point(pth))[0(𝕀)] a ∈ {X ⊢ _:A})

Lemma: path-point-1
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[a,b:{X ⊢ _:A}]. ∀[pth:{X ⊢ _:(Path_A b)}].  ((path-point(pth))[1(𝕀)] b ∈ {X ⊢ _:A})

Definition: path-contraction
path-contraction(X;pth) ==  <>(((pth)p)p (q)p ∧ q)

Lemma: path-contraction_wf
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[a,b:{X ⊢ _:A}]. ∀[pth:{X ⊢ _:(Path_A b)}].
  (path-contraction(X;pth) ∈ {X.𝕀 ⊢ _:(Path_(A)p (a)p path-point(pth))})

Lemma: path-contraction-0
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[a,b:{X ⊢ _:A}]. ∀[pth:{X ⊢ _:(Path_A b)}].
  ((path-contraction(X;pth))[0(𝕀)] refl(a) ∈ {X ⊢ _:(Path_A a)})

Lemma: path-contraction-1
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[a,b:{X ⊢ _:A}]. ∀[pth:{X ⊢ _:(Path_A b)}].
  ((path-contraction(X;pth))[1(𝕀)] pth ∈ {X ⊢ _:(Path_A b)})

Definition: singleton-contraction
singleton-contraction(X;pth) ==  <>(cubical-pair((pth)p q;path-contraction(X;pth)))

Lemma: singleton-contraction_wf
[X:j⊢]. ∀[T:{X ⊢ _}]. ∀[a,b:{X ⊢ _:T}]. ∀[pth:{X ⊢ _:(Path_T b)}].
  (singleton-contraction(X;pth) ∈ {X ⊢ _:(Path_Σ (Path_(T)p (a)p q) cubical-pair(a;refl(a)) cubical-pair(b;pth))})

Definition: singleton-type
Singleton(a) ==  Σ (Path_(A)p (a)p q)

Lemma: singleton-type_wf
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[a:{X ⊢ _:A}].  X ⊢ Singleton(a)

Definition: singleton-center
singleton-center(X;a) ==  cubical-pair(a;refl(a))

Lemma: singleton-center_wf
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[a:{X ⊢ _:A}].  (singleton-center(X;a) ∈ {X ⊢ _:Singleton(a)})

Lemma: csm-singleton-center
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[a:{X ⊢ _:A}]. ∀[H:j⊢]. ∀[s:H j⟶ X].
  ((singleton-center(X;a))s singleton-center(H;(a)s) ∈ {H ⊢ _:Singleton((a)s)})

Lemma: singleton-contraction_wf2
[X:j⊢]. ∀[T:{X ⊢ _}]. ∀[a:{X ⊢ _:T}]. ∀[s:{X ⊢ _:Singleton(a)}].
  (singleton-contraction(X;s.2) ∈ {X ⊢ _:(Path_Singleton(a) singleton-center(X;a) s)})

Lemma: csm-singleton-type
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[a:{X ⊢ _:A}]. ∀[H:j⊢]. ∀[s:H j⟶ X].  ((Singleton(a))s Singleton((a)s) ∈ {H ⊢ _})

Definition: singleton-contr
singleton-contr(X;A;a) ==  contr-witness(X;singleton-center(X;a);singleton-contraction(X.Singleton(a);q.2))

Lemma: singleton-contr_wf
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[a:{X ⊢ _:A}].  (singleton-contr(X;A;a) ∈ {X ⊢ _:Contractible(Singleton(a))})

Lemma: contractibilty-of-singleton
X:j⊢. ∀A:{X ⊢ _}. ∀a:{X ⊢ _:A}.  {X ⊢ _:Contractible(Singleton(a))}

Definition: refl-contraction
refl-contraction(X;A;a) ==  singleton-contraction(X;refl(a))

Definition: id-fiber-center
id-fiber-center(X;T) ==  cubical-pair(q;refl(q))

Lemma: id-fiber-center_wf
[X:j⊢]. ∀[T:{X ⊢ _}].  (id-fiber-center(X;T) ∈ {X.T ⊢ _:Σ (T)p (Path_((T)p)p (q)p q)})

Lemma: csm-id-fiber-center
[G,K:j⊢]. ∀[tau:K j⟶ G]. ∀[A:{G ⊢ _}].
  (id-fiber-center(K;(A)tau) (id-fiber-center(G;A))tau+ ∈ {K.(A)tau ⊢ _:Fiber((cubical-id-fun(K))p;q)})

Definition: id-fiber-contraction
id-fiber-contraction(X;T) ==  (singleton-contraction(X.T.(T)p.(Path_((T)p)p (q)p q);q))SigmaElim

Lemma: id-fiber-contraction_wf
[X:j⊢]. ∀[T:{X ⊢ _}].
  (id-fiber-contraction(X;T) ∈ {X.T.Σ (T)p (Path_((T)p)p (q)p q) ⊢ _
                                :(Path_(Σ (T)p (Path_((T)p)p (q)p q))p (id-fiber-center(X;T))p q)})

Lemma: csm-id-fiber-contraction
[G,K:j⊢]. ∀[tau:K j⟶ G]. ∀[A:{G ⊢ _}].
  (id-fiber-contraction(K;(A)tau)
  (id-fiber-contraction(G;A))tau++
  ∈ {K.(A)tau.Σ ((A)tau)p (Path_(((A)tau)p)p (q)p q) ⊢ _
     :(Path_(Σ ((A)tau)p (Path_(((A)tau)p)p (q)p q))p (id-fiber-center(K;(A)tau))p q)})

Definition: cubical-id-is-equiv
cubical-id-is-equiv(X;T) ==  contr-witness(X.T;id-fiber-center(X;T);id-fiber-contraction(X;T)))

Lemma: cubical-id-is-equiv_wf
X:j⊢. ∀T:{X ⊢ _}.  (cubical-id-is-equiv(X;T) ∈ {X ⊢ _:IsEquiv(T;T;cubical-id-fun(X))})

Lemma: csm-cubical-id-is-equiv
[G,K:j⊢]. ∀[tau:K j⟶ G]. ∀[A:{G ⊢ _}].
  (cubical-id-is-equiv(K;(A)tau) (cubical-id-is-equiv(G;A))tau ∈ {K ⊢ _:IsEquiv((A)tau;(A)tau;cubical-id-fun(K))})

Definition: cubical-id-equiv
IdEquiv(X;T) ==  equiv-witness(cubical-id-fun(X);cubical-id-is-equiv(X;T))

Lemma: cubical-id-equiv_wf
X:j⊢. ∀T:{X ⊢ _}.  (IdEquiv(X;T) ∈ {X ⊢ _:Equiv(T;T)})

Lemma: csm-cubical-id-equiv
[G,K:j⊢]. ∀[tau:K j⟶ G]. ∀[A:{G ⊢ _}].  ((IdEquiv(G;A))tau IdEquiv(K;(A)tau) ∈ {K ⊢ _:Equiv((A)tau;(A)tau)})

Lemma: cubical-id-equiv-subset
[G,phi,A:Top].  (IdEquiv(G, phi;A) IdEquiv(G;A))

Lemma: equiv-fun-cubical-id-equiv
[G,A:Top].  (equiv-fun(IdEquiv(G;A)) cubical-id-fun(G))

Lemma: unit-pathtype
[G:j⊢]. ∀a,b:{G ⊢ _:Path(1)}.  (a b ∈ {G ⊢ _:Path(1)})

Lemma: unit-path-type
[G:j⊢]. ∀[x,y:{G ⊢ _:1}].  ∀a,b:{G ⊢ _:(Path_1 y)}.  (a b ∈ {G ⊢ _:(Path_1 y)})

Definition: h-level
h-level(X;n;A) ==  primrec(n;λA.Contractible(A);λi,r,A. (r ΠA Π(A)p (Path_((A)p)p (q)p q))) A

Lemma: h-level_wf
[n:ℕ]. ∀[X:j⊢]. ∀[A:{X ⊢ _}].  X ⊢ h-level(X;n;A)

Definition: is-prop
isProp(A) ==  ΠA Π(A)p (Path_((A)p)p (q)p q)

Lemma: is-prop_wf
[X:j⊢]. ∀[A:{X ⊢ _}].  X ⊢ isProp(A)

Lemma: is-prop-contractible
X:j⊢. ∀A:{X ⊢ _}.  ({X ⊢ _:isProp(A)}  {X ⊢ _:A}  {X ⊢ _:Contractible(A)})

Definition: cubical-fun-ext
cubical-fun-ext(X;e) ==  <>((λapp(((e)p)p; q) (q)p))

Lemma: cubical-fun-ext_wf
X:j⊢. ∀A:{X ⊢ _}. ∀B:{X.A ⊢ _}. ∀f,g:{X ⊢ _:ΠB}. ∀e:{X ⊢ _:Π(Path_B app((f)p; q) app((g)p; q))}.
  (cubical-fun-ext(X;e) ∈ {X ⊢ _:(Path_Πg)})

Lemma: discrete-path-endpoints
[X:j⊢]. ∀[T:Type].
  ∀a:{X ⊢ _:discr(T)}. ∀[b:{X ⊢ _:discr(T)}]. ∀[p:{X ⊢ _:(Path_discr(T) b)}].  (a b ∈ {X ⊢ _:discr(T)})

Lemma: discrete-pathtype
[T:Type]. ∀[X:j⊢]. ∀[pth:{X ⊢ _:Path(discr(T))}].  (pth refl(pth 0(𝕀)) ∈ {X ⊢ _:Path(discr(T))})

Lemma: discrete-pathtype-uniform
[T:Type]. ∀[X,Z:j⊢]. ∀[s:Z j⟶ X]. ∀[pth:{Z ⊢ _:(Path(discr(T)))s}].
  (pth refl(pth 0(𝕀)) ∈ {Z ⊢ _:(Path(discr(T)))s})

Lemma: discrete-path
[T:Type]. ∀[X:j⊢]. ∀[a,b:{X ⊢ _:discr(T)}]. ∀[p:{X ⊢ _:(Path_discr(T) b)}].
  (p refl(a) ∈ {X ⊢ _:(Path_discr(T) b)})

Lemma: discrete-path-equal
[T:Type]. ∀[X:j⊢]. ∀[a,b:{X ⊢ _:discr(T)}]. ∀[p1,p2:{X ⊢ _:(Path_discr(T) b)}].
  (p1 p2 ∈ {X ⊢ _:(Path_discr(T) b)})

Lemma: equal-fiber-when-discrete
[X:j⊢]. ∀[T,A:{X ⊢ _}]. ∀[f:{X ⊢ _:(A ⟶ T)}]. ∀[z:{X ⊢ _:T}]. ∀[a,b:{X ⊢ _:Fiber(f;z)}].
  b ∈ {X ⊢ _:Fiber(f;z)} ⇐⇒ a.1 b.1 ∈ {X ⊢ _:A} 
  supposing ∀x:{X ⊢ _:Path(T)}. (x refl(x 0(𝕀)) ∈ {X ⊢ _:Path(T)})

Lemma: equal-fiber-discrete
[B:Type]. ∀[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[f:{X ⊢ _:(A ⟶ discr(B))}]. ∀[z:{X ⊢ _:discr(B)}]. ∀[a,b:{X ⊢ _:Fiber(f;z)}].
  (a b ∈ {X ⊢ _:Fiber(f;z)} ⇐⇒ a.1 b.1 ∈ {X ⊢ _:A})

Lemma: fiber-discrete-equal
[B:Type]. ∀[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[f:{X ⊢ _:(A ⟶ discr(B))}]. ∀[z:{X ⊢ _:discr(B)}]. ∀[fbr:{X ⊢ _:Fiber(f;z)}].
  (app(f; fiber-member(fbr)) z ∈ {X ⊢ _:discr(B)})

Definition: discrete-fun
discrete-fun(f) ==  λI,alpha. (f alpha 1)

Lemma: discrete-fun_wf
[A,B:Type]. ∀[X:j⊢]. ∀[f:{X ⊢ _:(discr(A) ⟶ discr(B))}].  (discrete-fun(f) ∈ {X ⊢ _:discr(A ⟶ B)})

Lemma: discrete-fun-at-app
[f,a,alpha,I:Top].  (app(f; discr(a))(alpha) discrete-fun(f)(alpha) a)

Lemma: discrete-fun-at
[A,B:Type]. ∀[f:{() ⊢ _:(discr(A) ⟶ discr(B))}]. ∀[I:fset(ℕ)]. ∀[a:()(I)].
  ((f a) J,h,u. (f {} ⋅ {} u)) ∈ (discr(A) ⟶ discr(B))(a))

Lemma: discrete-fun-invariant
[A,B:Type]. ∀[f:{() ⊢ _:(discr(A) ⟶ discr(B))}]. ∀[I:fset(ℕ)]. ∀[a:()(I)].
  ((f a) (f {} ⋅) ∈ (discr(A) ⟶ discr(B))(a))

Lemma: discrete-fun-app-invariant
[A,B:Type]. ∀[f:{() ⊢ _:(discr(A) ⟶ discr(B))}]. ∀[I:fset(ℕ)]. ∀[a:()(I)]. ∀[t:A].
  (app(f; discr(t))(a) app(f; discr(t))(⋅) ∈ B)

Lemma: discrete-fun-bijection
[A,B:Type].  Bij({() ⊢ _:(discr(A) ⟶ discr(B))};{() ⊢ _:discr(A ⟶ B)};λf.discrete-fun(f))

Definition: discrete-family
discrete-family(A;a.B[a]) ==  <λI,alpha. B[snd(alpha)], λI,J,f,alpha,x. x>

Lemma: discrete-family_wf
[A:Type]. ∀[B:A ⟶ Type]. ∀[X:j⊢].  X.discr(A) ⊢ discrete-family(A;a.B[a])

Lemma: csm-discrete-family
[A,B,s:Top].  ((discrete-family(A;a.B[a]))(s;q) discrete-family(A;a.B[a]))

Lemma: csm-discrete-pi
[A:Type]. ∀[B:A ⟶ Type]. ∀[X,Y:j⊢]. ∀[s:Y j⟶ X].
  ((Πdiscr(A) discrete-family(A;a.B[a]))s Y ⊢ Πdiscr(A) discrete-family(A;a.B[a]) ∈ {Y ⊢ _})

Lemma: csm-discrete-sigma
[A:Type]. ∀[B:A ⟶ Type]. ∀[X,Y:j⊢]. ∀[s:Y j⟶ X].
  ((Σ discr(A) discrete-family(A;a.B[a]))s = Σ discr(A) discrete-family(A;a.B[a]) ∈ {Y ⊢ _})

Definition: discrete-function
discrete-function(f) ==  λI,alpha. (f alpha 1)

Lemma: discrete-function_wf
[A:Type]. ∀[B:A ⟶ Type]. ∀[X:j⊢]. ∀[f:{X ⊢ _:Πdiscr(A) discrete-family(A;a.B[a])}].
  (discrete-function(f) ∈ {X ⊢ _:discr(a:A ⟶ B[a])})

Definition: discrete-function-inv
discrete-function-inv(X; b) ==  (λλI,alpha. (b(fst(alpha)) q(alpha)))

Lemma: discrete-function-inv_wf
[A:Type]. ∀[B:A ⟶ Type]. ∀[X:j⊢]. ∀[b:{X ⊢ _:discr(a:A ⟶ B[a])}].
  (discrete-function-inv(X; b) ∈ {X ⊢ _:Πdiscr(A) discrete-family(A;a.B[a])})

Lemma: discrete-function-inv-property
[A:Type]. ∀[B:A ⟶ Type]. ∀[X:j⊢]. ∀[b:{X ⊢ _:discr(a:A ⟶ B[a])}].
  (discrete-function(discrete-function-inv(X; b)) b ∈ {X ⊢ _:discr(a:A ⟶ B[a])})

Lemma: discrete-function-injection
[A:Type]. ∀[B:A ⟶ Type]. ∀[X:j⊢].
  ∀f,g:{X ⊢ _:Πdiscr(A) discrete-family(A;a.B[a])}.
    g ∈ {X ⊢ _:Πdiscr(A) discrete-family(A;a.B[a])} 
    supposing discrete-function(f) discrete-function(g) ∈ {X ⊢ _:discr(a:A ⟶ B[a])}

Lemma: discrete-pi-bijection
[A:Type]. ∀[B:A ⟶ Type].  ∀X:j⊢{X ⊢ _:Πdiscr(A) discrete-family(A;a.B[a])} {X ⊢ _:discr(a:A ⟶ B[a])}

Lemma: discrete-pi-equiv
A:Type. ∀B:A ⟶ Type. ∀X:j⊢.  {X ⊢ _:Equiv(Πdiscr(A) discrete-family(A;a.B[a]);discr(a:A ⟶ B[a]))}

Definition: discrete-pair
discrete-pair(p) ==  λI,alpha. <p.1(alpha), p.2(alpha)>

Lemma: discrete-pair_wf
[A:Type]. ∀[B:A ⟶ Type]. ∀[X:j⊢]. ∀[p:{X ⊢ _:Σ discr(A) discrete-family(A;a.B[a])}].
  (discrete-pair(p) ∈ {X ⊢ _:discr(a:A × B[a])})

Definition: discrete-pair-inv
discrete-pair-inv(X;b) ==  cubical-pair(λI,alpha. (fst(b(alpha)));λI,alpha. (snd(b(alpha))))

Lemma: discrete-pair-inv_wf
[A:Type]. ∀[B:A ⟶ Type]. ∀[X:j⊢]. ∀[b:{X ⊢ _:discr(a:A × B[a])}].
  (discrete-pair-inv(X;b) ∈ {X ⊢ _:Σ discr(A) discrete-family(A;a.B[a])})

Lemma: discrete-pair-inv-property
[A:Type]. ∀[B:A ⟶ Type]. ∀[X:j⊢]. ∀[b:{X ⊢ _:discr(a:A × B[a])}].
  (discrete-pair(discrete-pair-inv(X;b)) b ∈ {X ⊢ _:discr(a:A × B[a])})

Lemma: discrete-pair-injection
[A:Type]. ∀[B:A ⟶ Type]. ∀[X:j⊢].
  ∀f,g:{X ⊢ _:Σ discr(A) discrete-family(A;a.B[a])}.
    g ∈ {X ⊢ _:Σ discr(A) discrete-family(A;a.B[a])} 
    supposing discrete-pair(f) discrete-pair(g) ∈ {X ⊢ _:discr(a:A × B[a])}

Lemma: discrete-sigma-bijection
[A:Type]. ∀[B:A ⟶ Type].  ∀X:j⊢{X ⊢ _:Σ discr(A) discrete-family(A;a.B[a])} {X ⊢ _:discr(a:A × B[a])}

Lemma: discrete-sigma-equiv
A:Type. ∀B:A ⟶ Type. ∀X:j⊢.  {X ⊢ _:Equiv(Σ discr(A) discrete-family(A;a.B[a]);discr(a:A × B[a]))}

Definition: bijection-equiv
bijection-equiv(X;A;B;f;g) ==
  equiv-witness(cubical-lam(X;λI,alpha. (f (snd(alpha))));
                contr-witness(X.discr(B);fiber-point(λI,alpha. (g (snd(alpha)));refl(q));refl(q))))

Lemma: bijection-equiv_wf
[A,B:Type]. ∀[f:A ⟶ B]. ∀[g:B ⟶ A].
  (∀[X:j⊢]. (bijection-equiv(X;A;B;f;g) ∈ {X ⊢ _:Equiv(discr(A);discr(B))})) supposing 
     ((∀a:A. ((g (f a)) a ∈ A)) and 
     (∀b:B. ((f (g b)) b ∈ B)))

Lemma: equiv-discrete-type
A,B:Type. ∀f:A ⟶ B.  (Bij(A;B;f)  (∀X:j⊢{X ⊢ _:Equiv(discr(A);discr(B))}))

Definition: equiv-bijection
equiv-bijection(e) ==  discrete-fun(equiv-fun(e))(⋅)

Lemma: equiv-bijection_wf
[A,B:Type]. ∀[e:{() ⊢ _:Equiv(discr(A);discr(B))}].  (equiv-bijection(e) ∈ A ⟶ B)

Definition: equiv-bijection-is
equiv-bijection-is(e) ==  <λa1,a2,eq. Ax, λb.<equiv-contr(e;discr(b)).1.1(⋅), Ax>>

Lemma: equiv-bijection-is_wf
[A,B:Type]. ∀[e:{() ⊢ _:Equiv(discr(A);discr(B))}].  (equiv-bijection-is(e) ∈ Bij(A;B;equiv-bijection(e)))

Lemma: equiv-bijection-property
A,B:Type. ∀e:{() ⊢ _:Equiv(discr(A);discr(B))}.  Bij(A;B;equiv-bijection(e))

Lemma: discrete-equiv-iff
A,B:Type.  ({() ⊢ _:Equiv(discr(A);discr(B))} ⇐⇒ B)

Lemma: equiv-bijection-equiv
[A,B:Type]. ∀[e:A B].
  ((((λe.<equiv-bijection(e), equiv-bijection-is(e)>e.bijection-equiv(();A;B;fst(e);bij_inv(snd(e))))) e)
  e
  ∈ B)

Lemma: paths-are-refl-iff
[X:j⊢]. ∀[A:{X ⊢ _}].
  uiff(∀Z:j⊢. ∀s:Z j⟶ X. ∀p:{Z ⊢ _:Path((A)s)}.  (p refl(p 0(𝕀)) ∈ {Z ⊢ _:Path((A)s)});∀Z:j⊢. ∀s:Z j⟶ X.
                                                                                            ∀p:{Z ⊢ _:Path((A)s)}.
                                                                                              ∀[x,y:{Z ⊢ _:𝕀}].
                                                                                                (p x
                                                                                                y
                                                                                                ∈ {Z ⊢ _:(A)s}))

Lemma: paths-are-refl-iff2
[X:j⊢]. ∀[A:{X ⊢ _}].
  uiff(∀Z:j⊢. ∀s:Z j⟶ X. ∀p:{Z ⊢ _:Path((A)s)}.  (p refl(p 0(𝕀)) ∈ {Z ⊢ _:Path((A)s)});∀Z:j⊢. ∀s:Z j⟶ X.
                                                                                            ∀p:{Z.𝕀 ⊢ _:((A)s)p}.
                                                                                              ∀[x,y:{Z ⊢ _:𝕀}].
                                                                                                ((p)[x]
                                                                                                (p)[y]
                                                                                                ∈ {Z ⊢ _:(A)s}))

Definition: refl-map
refl-map(X;A) ==  cubical-lam(X;discr(refl(q)))

Lemma: refl-map_wf
[X:j⊢]. ∀[A:{X ⊢ _}].  (refl-map(X;A) ∈ {X ⊢ _:(A ⟶ discr({X.A ⊢ _:Path((A)p)}))})

Definition: refl-inv
refl-inv(b) ==  λI,alpha. (b(alpha) 0(𝕀))p(fst(alpha))

Comment: composition structure doc
=========================== composition structure  =============================

In this section we define composition and filling operators for
cubical types.
composition-op
filling-op.
We prove that fill_from_comp constructs filling operator from
composition operator.
(Proved in the lemmas fill_from_comp_wf1fill_from_comp_property1,
fill_from_comp_wf2and fill_from_comp_wf.)

One main result is that pi-comp constructs composition operator
for ΠB  from composition operators for and B.
This is proved in the series of lemmas
pi-comp_wf1pi-comp_wf2pi-comp_wf3,
pi-comp-propertyand pi-comp-uniformityand finally pi-comp_wf.


Comment: definitions of composition
Coquand's original formulation of composition structure
for type ⌜Gamma ⊢ A⌝ is defined in composition-op.
The type, Gamma ⊢ CompOp(A), is the type of composition operations.
Thought of as proposition, it says
 
⌜∀I:fset(ℕ). ∀i:{i:ℕ| ¬i ∈ I} . ∀rho:Gamma(I+i). ∀phi:𝔽(I). ∀u:{I+i,s(phi) ⊢ _:(A)<rho> iota}. ∀a0:A((i0)(rho)).
   ((∀J:fset(ℕ). ∀f:I,phi(J).  ((a0 (i0)(rho) f) u((i0) ⋅ f)))
    (∃a1:A((i1)(rho)). ∀J:fset(ℕ). ∀f:I,phi(J).  ((a1 (i1)(rho) f) u((i1) ⋅ f))))⌝
In addition, the operation, comp,  that builds a1 from the other inputs
must satisfy composition-uniformity(Gamma;A;comp)

This definition quantifies only over representable cubical sets of the form
I+i,s(phi), and also over the sets rho:Gamma(I+i).
So, when Gamma is cubical set in universe{n} then this definition of
composition structure is also type in universe{n}.

uniform composition operation of this kind can be used to build
cubical term composition-term that satisfies this typing rule 
composition-term_wf

The type of such construction is composition-function
and it is uniform if uniform-comp-function.

This gives another definition of composition structure:
 composition-structure
and the two definitions are equivalent:
composition-op-implies-composition-structure
composition-structure-implies-composition-op

The second definition quantifies over all cubical sets, so when Gamma
is cubical set in universe{n}, then composition structure for Gamma
is proposition in universe{n+1}. Nevertheless, the two definitions are equivalent.

We use the composition-op version to show composition for ΠB, Σ B, and (Path_A b):
pi-comp_wf
sigmacomp_wf
path_comp_wf

We will use the second definition to build composition structure for 
Glue [phi ⊢→ (T;w)] A

But, when we define the universes of "fibrant" types we use the original 
Gamma ⊢ CompOp(A) definition in order to get cumulativity
(viz. that term in ⌜c𝕌{n}⌝ is also in ⌜c𝕌{n+1}⌝).



Lemma: comp-nc-0-subset-I_cube
[I:fset(ℕ)]. ∀[i:{i:ℕ| ¬i ∈ I} ]. ∀[phi:𝔽(I)].  ∀J:fset(ℕ). ∀[f:I,phi(J)]. ((i0) ⋅ f ∈ I+i,s(phi)(J))

Lemma: comp-nc-1-subset-I_cube
[I:fset(ℕ)]. ∀[i:{i:ℕ| ¬i ∈ I} ]. ∀[phi:𝔽(I)].  ∀J:fset(ℕ). ∀[f:I,phi(J)]. ((i1) ⋅ f ∈ I+i,s(phi)(J))

Lemma: csm-cubical-subset-type-lemma
[Gamma,Delta:j⊢]. ∀[sigma:Delta j⟶ Gamma]. ∀[A:{Gamma ⊢ _}]. ∀[I:fset(ℕ)]. ∀[i:{i:ℕ| ¬i ∈ I} ]. ∀[rho:Delta(I+i)].
[J:fset(ℕ)]. ∀[f:J ⟶ I].
  (((A)sigma(f((i1)(rho))) A(f((i1)((sigma)rho))) ∈ Type) ∧ ((A)sigma(f((i0)(rho))) A(f((i0)((sigma)rho))) ∈ Type))

Definition: cubical-path-condition
cubical-path-condition(Gamma;A;I;i;rho;phi;u;a0) ==
  ∀J:fset(ℕ). ∀f:I,phi(J).  ((a0 (i0)(rho) f) u((i0) ⋅ f) ∈ A(f((i0)(rho))))

Lemma: cubical-path-condition_wf
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[I:fset(ℕ)]. ∀[i:{i:ℕ| ¬i ∈ I} ]. ∀[rho:Gamma(I+i)]. ∀[phi:𝔽(I)].
[u:{I+i,s(phi) ⊢ _:(A)<rho> iota}]. ∀[a0:A((i0)(rho))].
  (cubical-path-condition(Gamma;A;I;i;rho;phi;u;a0) ∈ ℙ')

Lemma: cubical-path-condition-0
[I:fset(ℕ)]. ∀[Gamma,A,i,rho,u,a0:Top].  cubical-path-condition(Gamma;A;I;i;rho;0;u;a0)

Lemma: sq_stable__cubical-path-condition
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[I:fset(ℕ)]. ∀[i:{i:ℕ| ¬i ∈ I} ]. ∀[rho:Gamma(I+i)]. ∀[phi:𝔽(I)].
[u:{I+i,s(phi) ⊢ _:(A)<rho> iota}]. ∀[a0:A((i0)(rho))].
  SqStable(cubical-path-condition(Gamma;A;I;i;rho;phi;u;a0))

Definition: cubical-path-condition'
cubical-path-condition'(Gamma;A;I;i;rho;phi;u;a1) ==
  ∀J:fset(ℕ). ∀f:I,phi(J).  ((a1 (i1)(rho) f) u((i1) ⋅ f) ∈ A(f((i1)(rho))))

Lemma: cubical-path-condition'_wf
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[I:fset(ℕ)]. ∀[i:{i:ℕ| ¬i ∈ I} ]. ∀[rho:Gamma(I+i)]. ∀[phi:𝔽(I)].
[u:{I+i,s(phi) ⊢ _:(A)<rho> iota}]. ∀[a1:A((i1)(rho))].
  (cubical-path-condition'(Gamma;A;I;i;rho;phi;u;a1) ∈ ℙ')

Lemma: sq_stable__cubical-path-condition'
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[I:fset(ℕ)]. ∀[i:{i:ℕ| ¬i ∈ I} ]. ∀[rho:Gamma(I+i)]. ∀[phi:𝔽(I)].
[u:{I+i,s(phi) ⊢ _:(A)<rho> iota}]. ∀[a0:A((i1)(rho))].
  SqStable(cubical-path-condition'(Gamma;A;I;i;rho;phi;u;a0))

Definition: cubical-path-0
cubical-path-0(Gamma;A;I;i;rho;phi;u) ==  {a0:A((i0)(rho))| cubical-path-condition(Gamma;A;I;i;rho;phi;u;a0)} 

Lemma: cubical-path-0_wf
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[I:fset(ℕ)]. ∀[i:{i:ℕ| ¬i ∈ I} ]. ∀[rho:Gamma(I+i)]. ∀[phi:𝔽(I)].
[u:{I+i,s(phi) ⊢ _:(A)<rho> iota}].
  (cubical-path-0(Gamma;A;I;i;rho;phi;u) ∈ 𝕌{[i' j']})

Lemma: cubical-path-0-0
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[I:fset(ℕ)]. ∀[i:{i:ℕ| ¬i ∈ I} ]. ∀[rho:Gamma(I+i)]. ∀[u:Top].
  cubical-path-0(Gamma;A;I;i;rho;0;u) ≡ A((i0)(rho))

Lemma: member-cubical-path-0-0
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[I:fset(ℕ)]. ∀[i:{i:ℕ| ¬i ∈ I} ]. ∀[rho:Gamma(I+i)]. ∀[u:Top]. ∀[x:A((i0)(rho))].
  (x ∈ cubical-path-0(Gamma;A;I;i;rho;0;u))

Definition: cubical-path-1
cubical-path-1(Gamma;A;I;i;rho;phi;u) ==  {a1:A((i1)(rho))| cubical-path-condition'(Gamma;A;I;i;rho;phi;u;a1)} 

Lemma: cubical-path-1_wf
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[I:fset(ℕ)]. ∀[i:{i:ℕ| ¬i ∈ I} ]. ∀[rho:Gamma(I+i)]. ∀[phi:𝔽(I)].
[u:{I+i,s(phi) ⊢ _:(A)<rho> iota}].
  (cubical-path-1(Gamma;A;I;i;rho;phi;u) ∈ 𝕌{[i' j']})

Lemma: cubical-path-1-ap-morph
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[I:fset(ℕ)]. ∀[i:{i:ℕ| ¬i ∈ I} ]. ∀[rho:Gamma(I+i)]. ∀[phi:𝔽(I)].
[u:{I+i,s(phi) ⊢ _:(A)<rho> iota}]. ∀[a:cubical-path-1(Gamma;A;I;i;rho;phi;u)]. ∀[J:fset(ℕ)]. ∀[g:J ⟶ I].
[j:{j:ℕ| ¬j ∈ J} ].
  ((a (i1)(rho) g) ∈ cubical-path-1(Gamma;A;J;j;g,i=j(rho);g(phi);(u)subset-trans(I+i;J+j;g,i=j;s(phi))))

Lemma: cubical-path-0-ap-morph
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[I:fset(ℕ)]. ∀[i:{i:ℕ| ¬i ∈ I} ]. ∀[rho:Gamma(I+i)]. ∀[phi:𝔽(I)].
[u:{I+i,s(phi) ⊢ _:(A)<rho> iota}]. ∀[a:cubical-path-0(Gamma;A;I;i;rho;phi;u)]. ∀[J:fset(ℕ)]. ∀[g:J ⟶ I].
[j:{j:ℕ| ¬j ∈ J} ].
  ((a (i0)(rho) g) ∈ cubical-path-0(Gamma;A;J;j;g,i=j(rho);g(phi);(u)subset-trans(I+i;J+j;g,i=j;s(phi))))

Lemma: cubical-subset-term-trans
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[I,J:fset(ℕ)]. ∀[i:{i:ℕ| ¬i ∈ I} ]. ∀[j:{j:ℕ| ¬j ∈ J} ]. ∀[g:J ⟶ I]. ∀[rho:Gamma(I+i)].
[phi:𝔽(I)]. ∀[u:{I+i,s(phi) ⊢ _:(A)<rho> iota}].
  ((u)subset-trans(I+i;J+j;g,i=j;s(phi)) ∈ {J+j,s(g(phi)) ⊢ _:(A)<g,i=j(rho)> iota})

Lemma: csm-cubical-path-0-subtype
[Gamma,Delta:j⊢]. ∀[sigma:Delta j⟶ Gamma]. ∀[A:{Gamma ⊢ _}]. ∀[I:fset(ℕ)]. ∀[i:{i:ℕ| ¬i ∈ I} ]. ∀[rho:Delta(I+i)].
[phi:𝔽(I)]. ∀[u:{I+i,s(phi) ⊢ _:((A)sigma)<rho> iota}].
  (cubical-path-0(Delta;(A)sigma;I;i;rho;phi;u) ⊆cubical-path-0(Gamma;A;I;i;(sigma)rho;phi;u))

Lemma: csm-cubical-path-0-subtype2
[Gamma,Delta:j⊢]. ∀[sigma:Delta j⟶ Gamma]. ∀[A:{Gamma ⊢ _}]. ∀[I:fset(ℕ)]. ∀[i:{i:ℕ| ¬i ∈ I} ]. ∀[rho:Delta(I+i)].
[phi:𝔽(I)]. ∀[u1,u2:{I+i,s(phi) ⊢ _:((A)sigma)<rho> iota}].
  cubical-path-0(Delta;(A)sigma;I;i;rho;phi;u1) ⊆cubical-path-0(Gamma;A;I;i;(sigma)rho;phi;u2) 
  supposing u1 u2 ∈ {I+i,s(phi) ⊢ _:((A)sigma)<rho> iota}

Lemma: csm-cubical-path-1-subtype
[Gamma,Delta:j⊢]. ∀[sigma:Delta j⟶ Gamma]. ∀[A:{Gamma ⊢ _}]. ∀[I:fset(ℕ)]. ∀[i:{i:ℕ| ¬i ∈ I} ]. ∀[rho:Delta(I+i)].
[phi:𝔽(I)]. ∀[u:{I+i,s(phi) ⊢ _:((A)sigma)<rho> iota}].
  (cubical-path-1(Gamma;A;I;i;(sigma)rho;phi;u) ⊆cubical-path-1(Delta;(A)sigma;I;i;rho;phi;u))

Definition: comp-op
comp-op(Gamma;A) ==
  I:fset(ℕ)
  ⟶ i:{i:ℕ| ¬i ∈ I} 
  ⟶ rho:Gamma(I+i)
  ⟶ phi:𝔽(I)
  ⟶ u:{I+i,s(phi) ⊢ _:(A)<rho> iota}
  ⟶ cubical-path-0(Gamma;A;I;i;rho;phi;u)
  ⟶ cubical-path-1(Gamma;A;I;i;rho;phi;u)

Lemma: comp-op_wf
[Gamma:j⊢]. ∀[A:{Gamma ⊢_}].  (comp-op(Gamma;A) ∈ 𝕌{[i j]'})

Definition: composition-uniformity
composition-uniformity(Gamma;A;comp) ==
  ∀I,J:fset(ℕ). ∀i:{i:ℕ| ¬i ∈ I} . ∀j:{j:ℕ| ¬j ∈ J} . ∀g:J ⟶ I. ∀rho:Gamma(I+i). ∀phi:𝔽(I).
  ∀u:{I+i,s(phi) ⊢ _:(A)<rho> iota}. ∀a0:cubical-path-0(Gamma;A;I;i;rho;phi;u).
    ((comp rho phi a0 (i1)(rho) g)
    (comp g,i=j(rho) g(phi) (u)subset-trans(I+i;J+j;g,i=j;s(phi)) (a0 (i0)(rho) g))
    ∈ A(g((i1)(rho))))

Lemma: composition-uniformity_wf
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[comp:I:fset(ℕ)
                                      ⟶ i:{i:ℕ| ¬i ∈ I} 
                                      ⟶ rho:Gamma(I+i)
                                      ⟶ phi:𝔽(I)
                                      ⟶ u:{I+i,s(phi) ⊢ _:(A)<rho> iota}
                                      ⟶ cubical-path-0(Gamma;A;I;i;rho;phi;u)
                                      ⟶ cubical-path-1(Gamma;A;I;i;rho;phi;u)].
  (composition-uniformity(Gamma;A;comp) ∈ ℙ{[i' j']})

Lemma: sq_stable__composition-uniformity
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[comp:I:fset(ℕ)
                                      ⟶ i:{i:ℕ| ¬i ∈ I} 
                                      ⟶ rho:Gamma(I+i)
                                      ⟶ phi:𝔽(I)
                                      ⟶ u:{I+i,s(phi) ⊢ _:(A)<rho> iota}
                                      ⟶ cubical-path-0(Gamma;A;I;i;rho;phi;u)
                                      ⟶ cubical-path-1(Gamma;A;I;i;rho;phi;u)].
  SqStable(composition-uniformity(Gamma;A;comp))

Definition: composition-op
Gamma ⊢ CompOp(A) ==
  {comp:I:fset(ℕ)
   ⟶ i:{i:ℕ| ¬i ∈ I} 
   ⟶ rho:Gamma(I+i)
   ⟶ phi:𝔽(I)
   ⟶ u:{I+i,s(phi) ⊢ _:(A)<rho> iota}
   ⟶ cubical-path-0(Gamma;A;I;i;rho;phi;u)
   ⟶ cubical-path-1(Gamma;A;I;i;rho;phi;u)| 
   composition-uniformity(Gamma;A;comp)} 

Lemma: composition-op_wf1
[Gamma:j⊢]. ∀[A:{Gamma ⊢_}].  (Gamma ⊢ CompOp(A) ∈ 𝕌{[i j]'})

Lemma: composition-op_wf
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}].  (Gamma ⊢ CompOp(A) ∈ 𝕌{[i' j']})

Lemma: equal-composition-op
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[c1:Gamma ⊢ CompOp(A)]. ∀[c2:I:fset(ℕ)
                                                             ⟶ i:{i:ℕ| ¬i ∈ I} 
                                                             ⟶ rho:Gamma(I+i)
                                                             ⟶ phi:𝔽(I)
                                                             ⟶ u:{I+i,s(phi) ⊢ _:(A)<rho> iota}
                                                             ⟶ cubical-path-0(Gamma;A;I;i;rho;phi;u)
                                                             ⟶ cubical-path-1(Gamma;A;I;i;rho;phi;u)].
  c1 c2 ∈ Gamma ⊢ CompOp(A) 
  supposing c1
  c2
  ∈ (I:fset(ℕ)
    ⟶ i:{i:ℕ| ¬i ∈ I} 
    ⟶ rho:Gamma(I+i)
    ⟶ phi:𝔽(I)
    ⟶ u:{I+i,s(phi) ⊢ _:(A)<rho> iota}
    ⟶ cubical-path-0(Gamma;A;I;i;rho;phi;u)
    ⟶ cubical-path-1(Gamma;A;I;i;rho;phi;u))

Lemma: equal-composition-op2
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[c1,c2:Gamma ⊢ CompOp(A)].
  c1 c2 ∈ Gamma ⊢ CompOp(A) 
  supposing ∀I:fset(ℕ). ∀i:{i:ℕ| ¬i ∈ I} . ∀rho:Gamma(I+i). ∀phi:𝔽(I). ∀u:{I+i,s(phi) ⊢ _:(A)<rho> iota}.
            ∀a0:cubical-path-0(Gamma;A;I;i;rho;phi;u).
              ((c1 rho phi a0) (c2 rho phi a0) ∈ A((i1)(rho)))

Lemma: composition-op-uniformity
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[comp:Gamma ⊢ CompOp(A)].
  ∀I,J:fset(ℕ). ∀i:{i:ℕ| ¬i ∈ I} . ∀j:{j:ℕ| ¬j ∈ J} . ∀g:J ⟶ I. ∀rho:Gamma(I+i). ∀phi:𝔽(I).
  ∀u:{I+i,s(phi) ⊢ _:(A)<rho> iota}. ∀a0:cubical-path-0(Gamma;A;I;i;rho;phi;u).
    ((comp rho phi a0 (i1)(rho) g)
    (comp g,i=j(rho) g(phi) (u)subset-trans(I+i;J+j;g,i=j;s(phi)) (a0 (i0)(rho) g))
    ∈ A(g((i1)(rho))))

Lemma: composition-op-nc-e
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[comp:Gamma ⊢ CompOp(A)].
  ∀I:fset(ℕ). ∀i,j:{j:ℕ| ¬j ∈ I} . ∀rho:Gamma(I+i). ∀phi:𝔽(I). ∀u:{I+i,s(phi) ⊢ _:(A)<rho> iota}.
  ∀a0:cubical-path-0(Gamma;A;I;i;rho;phi;u).
    ((comp rho phi a0) (comp e(i;j)(rho) phi (u)subset-trans(I+i;I+j;e(i;j);s(phi)) a0) ∈ A((i1)(rho)))

Definition: csm-composition
(comp)sigma ==  λI,i,rho. (comp (sigma)rho)

Lemma: csm-composition_wf
[Gamma,Delta:j⊢]. ∀[sigma:Delta j⟶ Gamma]. ∀[A:{Gamma ⊢ _}]. ∀[comp:Gamma ⊢ CompOp(A)].
  ((comp)sigma ∈ Delta ⊢ CompOp((A)sigma))

Lemma: csm-composition-exists
[Gamma,Delta:j⊢]. ∀[A:{Gamma ⊢ _}].  ∀s:Delta j⟶ Gamma. (Gamma ⊢ CompOp(A)  Delta ⊢ CompOp((A)s))

Lemma: csm-p-composition-exists
[X:j⊢]. ∀[A,T:{X ⊢ _}].  (X ⊢ CompOp(A)  X.T ⊢ CompOp((A)p))

Lemma: csm-composition-id
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[comp:Gamma ⊢ CompOp(A)].  ((comp)1(Gamma) comp ∈ Gamma ⊢ CompOp(A))

Lemma: csm-composition-comp
[X,Y,Z:j⊢]. ∀[s1:Z j⟶ Y]. ∀[s2:Y j⟶ X]. ∀[A:{X ⊢ _}]. ∀[comp:X ⊢ CompOp(A)].
  (((comp)s2)s1 (comp)s2 s1 ∈ Z ⊢ CompOp((A)s2 s1))

Lemma: composition-op-1
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[cA:Gamma ⊢ CompOp(A)]. ∀[I:fset(ℕ)]. ∀[i:{i:ℕ| ¬i ∈ I} ]. ∀[rho:Gamma(I+i)].
[u:{I+i,s(1) ⊢ _:(A)<rho> iota}]. ∀[a:cubical-path-0(Gamma;A;I;i;rho;1;u)]. ∀[J:fset(ℕ)]. ∀[f:J ⟶ I].
  ((cA rho (i1)(rho) f) u((i1) ⋅ f) ∈ A(f((i1)(rho))))

Lemma: composition-op-1-case1
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[cA:Gamma ⊢ CompOp(A)]. ∀[I:fset(ℕ)]. ∀[i:{i:ℕ| ¬i ∈ I} ]. ∀[rho:Gamma(I+i)].
[u:{I+i,s(1) ⊢ _:(A)<rho> iota}]. ∀[a:cubical-path-0(Gamma;A;I;i;rho;1;u)].
  ((cA rho a) u((i1)) ∈ A((i1)(rho)))

Definition: filling-uniformity
filling-uniformity(Gamma;A;fill) ==
  ∀I,J:fset(ℕ). ∀i:{i:ℕ| ¬i ∈ I} . ∀j:{j:ℕ| ¬j ∈ J} . ∀g:J ⟶ I. ∀rho:Gamma(I+i). ∀phi:𝔽(I).
  ∀u:{I+i,s(phi) ⊢ _:(A)<rho> iota}. ∀a0:cubical-path-0(Gamma;A;I;i;rho;phi;u).
    ((fill rho phi a0 rho g,i=j)
    (fill g,i=j(rho) g(phi) (u)subset-trans(I+i;J+j;g,i=j;s(phi)) (a0 (i0)(rho) g))
    ∈ A(g,i=j(rho)))

Lemma: filling-uniformity_wf
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[fill:I:fset(ℕ)
                                      ⟶ i:{i:ℕ| ¬i ∈ I} 
                                      ⟶ rho:Gamma(I+i)
                                      ⟶ phi:𝔽(I)
                                      ⟶ u:{I+i,s(phi) ⊢ _:(A)<rho> iota}
                                      ⟶ cubical-path-0(Gamma;A;I;i;rho;phi;u)
                                      ⟶ A(rho)].
  (filling-uniformity(Gamma;A;fill) ∈ ℙ{[i' j']})

Definition: filling-op
filling-op(Gamma;A) ==
  {fill:I:fset(ℕ)
   ⟶ i:{i:ℕ| ¬i ∈ I} 
   ⟶ rho:Gamma(I+i)
   ⟶ phi:𝔽(I)
   ⟶ u:{I+i,s(phi) ⊢ _:(A)<rho> iota}
   ⟶ a0:cubical-path-0(Gamma;A;I;i;rho;phi;u)
   ⟶ {a:A(rho)| 
       (section-iota(Gamma;A;I+i;rho;a) u ∈ {I+i,s(phi) ⊢ _:(A)<rho> iota}) ∧ ((a rho (i0)) a0 ∈ A((i0)(rho)))} 
   filling-uniformity(Gamma;A;fill)} 

Lemma: filling-op_wf
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}].  (filling-op(Gamma;A) ∈ 𝕌{[i' j']})

Definition: fillterm
fillterm(Gamma;A;I;i;j;rho;a0;u) ==  λK,f. if isdM0(f i) then (a0 (i0)(rho) s ⋅ f) else u(m(i;j) ⋅ f) fi 

Lemma: fillterm_wf
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[I:fset(ℕ)]. ∀[i:{i:ℕ| ¬i ∈ I} ]. ∀[j:{j:ℕ| ¬j ∈ I+i} ]. ∀[rho:Gamma(I+i)]. ∀[phi:𝔽(I)].
[u:{I+i,s(phi) ⊢ _:(A)<rho> iota}]. ∀[a0:cubical-path-0(Gamma;A;I;i;rho;phi;u)].
  (fillterm(Gamma;A;I;i;j;rho;a0;u) ∈ {I+i+j,s(fl-join(I+i;s(phi);(i=0))) ⊢ _:(A)<m(i;j)(rho)> iota})

Lemma: cubical-path-0-fillterm
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[I:fset(ℕ)]. ∀[i:{i:ℕ| ¬i ∈ I} ]. ∀[j:{j:ℕ| ¬j ∈ I+i} ]. ∀[rho:Gamma(I+i)]. ∀[phi:𝔽(I)].
  ∀u:{I+i,s(phi) ⊢ _:(A)<rho> iota}
    ∀[a0:cubical-path-0(Gamma;A;I;i;rho;phi;u)]
      ((a0 (i0)(rho) s)
       ∈ cubical-path-0(Gamma;A;I+i;j;m(i;j)(rho);fl-join(I+i;s(phi);(i=0));fillterm(Gamma;A;I;i;j;rho;a0;u)))

Definition: fill_from_comp
fill_from_comp(Gamma;A;comp) ==
  λI,i,rho,phi,u,a0. eval new-name(I+i) in
                     comp I+i m(i;j)(rho) fl-join(I+i;s(phi);(i=0)) fillterm(Gamma;A;I;i;j;rho;a0;u) (a0 (i0)(rho) s)

Lemma: fill_from_comp_wf1
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[comp:Gamma ⊢ CompOp(A)].
  (fill_from_comp(Gamma;A;comp) ∈ I:fset(ℕ)
   ⟶ i:{i:ℕ| ¬i ∈ I} 
   ⟶ rho:Gamma(I+i)
   ⟶ phi:𝔽(I)
   ⟶ u:{I+i,s(phi) ⊢ _:(A)<rho> iota}
   ⟶ a0:cubical-path-0(Gamma;A;I;i;rho;phi;u)
   ⟶ let new-name(I+i) in
          cubical-path-1(Gamma;A;I+i;j;m(i;j)(rho);fl-join(I+i;s(phi);(i=0));fillterm(Gamma;A;I;i;j;rho;a0;u)))

Lemma: fill_from_comp_property1
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[comp:Gamma ⊢ CompOp(A)]. ∀[I:fset(ℕ)]. ∀[i:{i:ℕ| ¬i ∈ I} ]. ∀[rho:Gamma(I+i)].
[phi:𝔽(I)]. ∀[u:{I+i,s(phi) ⊢ _:(A)<rho> iota}]. ∀[a0:cubical-path-0(Gamma;A;I;i;rho;phi;u)].
  ((fill_from_comp(Gamma;A;comp) rho phi a0 rho (i0)) a0 ∈ A((i0)(rho)))

Lemma: fill_from_comp_wf2
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[comp:Gamma ⊢ CompOp(A)].
  (fill_from_comp(Gamma;A;comp) ∈ I:fset(ℕ)
   ⟶ i:{i:ℕ| ¬i ∈ I} 
   ⟶ rho:Gamma(I+i)
   ⟶ phi:𝔽(I)
   ⟶ u:{I+i,s(phi) ⊢ _:(A)<rho> iota}
   ⟶ a0:cubical-path-0(Gamma;A;I;i;rho;phi;u)
   ⟶ {a:A(rho)| 
       (section-iota(Gamma;A;I+i;rho;a) u ∈ {I+i,s(phi) ⊢ _:(A)<rho> iota}) ∧ ((a rho (i0)) a0 ∈ A((i0)(rho)))} )

Lemma: fill_from_comp_wf
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[comp:Gamma ⊢ CompOp(A)].  (fill_from_comp(Gamma;A;comp) ∈ filling-op(Gamma;A))

Definition: pi-comp-nu
pi-comp-nu(Gamma;A;cA;I;i;rho;J;f;u1;j) ==
  let v' fill_from_comp(Gamma;A;cA) f,i=1-j(rho) () u1 in
      (v' f,i=1-j(rho) r_j)

Lemma: pi-comp-nu_wf
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[cA:Gamma ⊢ CompOp(A)]. ∀[I:fset(ℕ)]. ∀[i:{i:ℕ| ¬i ∈ I} ]. ∀[rho:Gamma(I+i)].
[J:fset(ℕ)]. ∀[f:J ⟶ I]. ∀[u1:A(f((i1)(rho)))]. ∀[j:{j:ℕ| ¬j ∈ J} ].
  (pi-comp-nu(Gamma;A;cA;I;i;rho;J;f;u1;j) ∈ A(r_j(f,i=1-j(rho))))

Lemma: pi-comp-nu-property
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[cA:Gamma ⊢ CompOp(A)]. ∀[I:fset(ℕ)]. ∀[i:{i:ℕ| ¬i ∈ I} ]. ∀[rho:Gamma(I+i)].
[J:fset(ℕ)]. ∀[f:J ⟶ I]. ∀[u1:A(f((i1)(rho)))]. ∀[j:{j:ℕ| ¬j ∈ J} ].
  ((pi-comp-nu(Gamma;A;cA;I;i;rho;J;f;u1;j) f,i=j(rho) (j1)) u1 ∈ A((j1)(f,i=j(rho))))

Lemma: pi-comp-nu-uniformity
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[cA:Gamma ⊢ CompOp(A)]. ∀[I:fset(ℕ)]. ∀[i:{i:ℕ| ¬i ∈ I} ]. ∀[rho:Gamma(I+i)].
[J,K:fset(ℕ)]. ∀[f:J ⟶ I]. ∀[g:K ⟶ J]. ∀[u:A(f((i1)(rho)))]. ∀[j:{j:ℕ| ¬j ∈ J} ]. ∀[k:{k:ℕ| ¬k ∈ K} ].
  ((pi-comp-nu(Gamma;A;cA;I;i;rho;J;f;u;j) f,i=j(rho) g,j=k)
  pi-comp-nu(Gamma;A;cA;I;i;rho;K;f ⋅ g;(u f((i1)(rho)) g);k)
  ∈ A(f ⋅ g,i=k(rho)))

Definition: pi-comp-app
pi-comp-app(Gamma;A;I;i;rho;phi;mu;J;f;j;nu) ==
  app((mu)subset-trans(I+i;J+j;f,i=j;s(phi)); (canonical-section(Gamma;A;J+j;f,i=j(rho);nu))iota)

Lemma: pi-comp-app_wf
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[B:{Gamma.A ⊢ _}]. ∀[I:fset(ℕ)]. ∀[i:{i:ℕ| ¬i ∈ I} ]. ∀[rho:Gamma(I+i)]. ∀[phi:𝔽(I)].
[mu:{I+i,s(phi) ⊢ _:(ΠB)<rho> iota}]. ∀[J:fset(ℕ)]. ∀[f:J ⟶ I]. ∀[j:{j:ℕ| ¬j ∈ J} ]. ∀[nu:A(r_j(f,i=1-j(rho)))].
  (pi-comp-app(Gamma;A;I;i;rho;phi;mu;J;f;j;nu) ∈ {J+j,s(f(phi)) ⊢ _:(B)<(f,i=j(rho);nu)> iota})

Definition: pi-comp-lambda
pi-comp-lambda(Gamma;A;I;i;rho;lambda;J;f;j;nu) ==  lambda (nu r_j(f,i=1-j(rho)) (j0))

Lemma: pi-comp-lambda_wf
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[B:{Gamma.A ⊢ _}]. ∀[I:fset(ℕ)]. ∀[i:{i:ℕ| ¬i ∈ I} ]. ∀[rho:Gamma(I+i)]. ∀[phi:𝔽(I)].
[mu:{I+i,s(phi) ⊢ _:(ΠB)<rho> iota}]. ∀[lambda:cubical-path-0(Gamma;ΠB;I;i;rho;phi;mu)]. ∀[J:fset(ℕ)].
[f:J ⟶ I]. ∀[j:{j:ℕ| ¬j ∈ J} ]. ∀[nu:A(r_j(f,i=1-j(rho)))].
  (pi-comp-lambda(Gamma;A;I;i;rho;lambda;J;f;j;nu)
   ∈ cubical-path-0(Gamma.A;B;J;j;(f,i=j(rho);nu);f(phi);pi-comp-app(Gamma;A;I;i;rho;phi;mu;J;f;j;nu)))

Definition: pi-comp
pi-comp(Gamma;A;B;cA;cB) ==
  λI,i,rho,phi,mu,lambda,J,f,u1. eval new-name(J) in
                                 let nu pi-comp-nu(Gamma;A;cA;I;i;rho;J;f;u1;j) in
                                     cB (f,i=j(rho);nu) f(phi) pi-comp-app(Gamma;A;I;i;rho;phi;mu;J;f;j;nu) 
                                     pi-comp-lambda(Gamma;A;I;i;rho;lambda;J;f;j;nu)

Lemma: pi-comp_wf1
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[B:{Gamma.A ⊢ _}]. ∀[cA:Gamma ⊢ CompOp(A)]. ∀[cB:Gamma.A ⊢ CompOp(B)].
  (pi-comp(Gamma;A;B;cA;cB) ∈ I:fset(ℕ)
   ⟶ i:{i:ℕ| ¬i ∈ I} 
   ⟶ rho:Gamma(I+i)
   ⟶ phi:𝔽(I)
   ⟶ mu:{I+i,s(phi) ⊢ _:(ΠB)<rho> iota}
   ⟶ lambda:cubical-path-0(Gamma;ΠB;I;i;rho;phi;mu)
   ⟶ J:fset(ℕ)
   ⟶ f:J ⟶ I
   ⟶ u1:A(f((i1)(rho)))
   ⟶ let new-name(J) in
       let nu pi-comp-nu(Gamma;A;cA;I;i;rho;J;f;u1;j) in
       cubical-path-1(Gamma.A;B;J;j;(f,i=j(rho);nu);f(phi);pi-comp-app(Gamma;A;I;i;rho;phi;mu;J;f;j;nu)))

Lemma: pi-comp_wf2
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[B:{Gamma.A ⊢ _}]. ∀[cA:Gamma ⊢ CompOp(A)]. ∀[cB:Gamma.A ⊢ CompOp(B)].
  (pi-comp(Gamma;A;B;cA;cB) ∈ I:fset(ℕ)
   ⟶ i:{i:ℕ| ¬i ∈ I} 
   ⟶ rho:Gamma(I+i)
   ⟶ phi:𝔽(I)
   ⟶ mu:{I+i,s(phi) ⊢ _:(ΠB)<rho> iota}
   ⟶ lambda:cubical-path-0(Gamma;ΠB;I;i;rho;phi;mu)
   ⟶ J:fset(ℕ)
   ⟶ f:J ⟶ I
   ⟶ u1:A(f((i1)(rho)))
   ⟶ B((f((i1)(rho));u1)))

Lemma: pi-comp_wf3
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[B:{Gamma.A ⊢ _}]. ∀[cA:Gamma ⊢ CompOp(A)]. ∀[cB:Gamma.A ⊢ CompOp(B)].
  (pi-comp(Gamma;A;B;cA;cB) ∈ I:fset(ℕ)
   ⟶ i:{i:ℕ| ¬i ∈ I} 
   ⟶ rho:Gamma(I+i)
   ⟶ phi:𝔽(I)
   ⟶ mu:{I+i,s(phi) ⊢ _:(ΠB)<rho> iota}
   ⟶ lambda:cubical-path-0(Gamma;ΠB;I;i;rho;phi;mu)
   ⟶ ΠB((i1)(rho)))

Lemma: pi-comp-property
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[B:{Gamma.A ⊢ _}]. ∀[cA:Gamma ⊢ CompOp(A)]. ∀[cB:Gamma.A ⊢ CompOp(B)]. ∀[I:fset(ℕ)].
[i:{i:ℕ| ¬i ∈ I} ]. ∀[rho:Gamma(I+i)]. ∀[phi:𝔽(I)]. ∀[mu:{I+i,s(phi) ⊢ _:(ΠB)<rho> iota}].
[lambda:cubical-path-0(Gamma;ΠB;I;i;rho;phi;mu)]. ∀[J:fset(ℕ)]. ∀[f:I,phi(J)].
  ((pi-comp(Gamma;A;B;cA;cB) rho phi mu lambda (i1)(rho) f) mu((i1) ⋅ f) ∈ ΠB(f((i1)(rho))))

Lemma: pi-comp-uniformity
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[B:{Gamma.A ⊢ _}]. ∀[cA:Gamma ⊢ CompOp(A)]. ∀[cB:Gamma.A ⊢ CompOp(B)].
  composition-uniformity(Gamma;ΠB;pi-comp(Gamma;A;B;cA;cB))

Lemma: pi-comp_wf
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[B:{Gamma.A ⊢ _}]. ∀[cA:Gamma ⊢ CompOp(A)]. ∀[cB:Gamma.A ⊢ CompOp(B)].
  (pi-comp(Gamma;A;B;cA;cB) ∈ Gamma ⊢ CompOp(ΠB))

Definition: fun-comp
fun-comp(Gamma; A; B; cA; cB) ==  pi-comp(Gamma;A;(B)p;cA;(cB)p)

Lemma: fun-comp_wf
[Gamma:j⊢]. ∀[A,B:{Gamma ⊢ _}]. ∀[cA:Gamma ⊢ CompOp(A)]. ∀[cB:Gamma ⊢ CompOp(B)].
  (fun-comp(Gamma; A; B; cA; cB) ∈ Gamma ⊢ CompOp((A ⟶ B)))

Lemma: fun-comp-exists
Gamma:j⊢. ∀A,B:{Gamma ⊢ _}.  (Gamma ⊢ CompOp(A)  Gamma ⊢ CompOp(B)  Gamma ⊢ CompOp((A ⟶ B)))

Definition: sigmacomp
sigmacomp(Gamma;A;B;cA;cB) ==
  λI,i,rho,phi,mu,lambda. let fill_from_comp(Gamma;A;cA) rho phi mu.1 (fst(lambda)) in
                           let cB (rho;v) phi mu.2 (snd(lambda)) in
                           <(v rho (i1)), w>

Lemma: sigmacomp_wf1
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[B:{Gamma.A ⊢ _}]. ∀[cA:Gamma ⊢ CompOp(A)]. ∀[cB:Gamma.A ⊢ CompOp(B)].
  (sigmacomp(Gamma;A;B;cA;cB) ∈ I:fset(ℕ)
   ⟶ i:{i:ℕ| ¬i ∈ I} 
   ⟶ rho:Gamma(I+i)
   ⟶ phi:𝔽(I)
   ⟶ mu:{I+i,s(phi) ⊢ _:(Σ B)<rho> iota}
   ⟶ lambda:cubical-path-0(Gamma;Σ B;I;i;rho;phi;mu)
   ⟶ cubical-path-1(Gamma;Σ B;I;i;rho;phi;mu))

Lemma: sigmacomp_wf
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[B:{Gamma.A ⊢ _}]. ∀[cA:Gamma ⊢ CompOp(A)]. ∀[cB:Gamma.A ⊢ CompOp(B)].
  (sigmacomp(Gamma;A;B;cA;cB) ∈ Gamma ⊢ CompOp(Σ B))

Lemma: sigma-comp-exists
Gamma:j⊢. ∀A:{Gamma ⊢ _}. ∀B:{Gamma.A ⊢ _}.  (Gamma ⊢ CompOp(A)  Gamma.A ⊢ CompOp(B)  Gamma ⊢ CompOp(Σ B))

Lemma: pi-comp-exists
Gamma:j⊢. ∀A:{Gamma ⊢ _}. ∀B:{Gamma.A ⊢ _}.  (Gamma ⊢ CompOp(A)  Gamma.A ⊢ CompOp(B)  Gamma ⊢ CompOp(ΠB))

Lemma: composition-type-lemma1
[Gamma:j⊢]. ∀[A:{Gamma.𝕀 ⊢ _}]. ∀[I:fset(ℕ)]. ∀[rho:Gamma(I)].
  ((A)[0(𝕀)](rho) A((new-name(I)0)((s(rho);<new-name(I)>))) ∈ Type)

Lemma: composition-type-lemma2
[Gamma:j⊢]. ∀[A:{Gamma.𝕀 ⊢ _}]. ∀[I:fset(ℕ)]. ∀[rho:Gamma(I)].
  (A((new-name(I)1)((s(rho);<new-name(I)>))) (A)[1(𝕀)](rho) ∈ Type)

Lemma: context-map-lemma1
[Gamma:j⊢]. ∀[phi:{Gamma ⊢ _:𝔽}]. ∀[I:fset(ℕ)]. ∀[rho:Gamma(I)]. ∀[i:ℕ].  (<s(rho)> ∈ I+i,s(phi(rho)) j⟶ Gamma, phi)

Lemma: context-map-lemma2
[Gamma:j⊢]. ∀[phi:{Gamma ⊢ _:𝔽}]. ∀[I:fset(ℕ)]. ∀[rho:Gamma(I)].
  (<(s(rho);<new-name(I)>)> ∈ I+new-name(I),s(phi(rho)) j⟶ Gamma, phi.𝕀)

Lemma: context-subset-ap-iota
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[phi:{Gamma ⊢ _:𝔽}].  ((A)iota A ∈ {Gamma, phi ⊢ _})

Lemma: subset-iota-is-id
[H:j⊢]. ∀[phi:{H ⊢ _:𝔽}]. ∀[v:{H, phi ⊢ _}].  (v (v)iota ∈ {H, phi ⊢ _})

Lemma: context-subset-term-iota
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[phi:{Gamma ⊢ _:𝔽}]. ∀[v:{Gamma, phi ⊢ _:A}].  ((v)iota v ∈ {Gamma, phi ⊢ _:A})

Lemma: context-subset-type-iota
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[phi:{Gamma ⊢ _:𝔽}].  ({Gamma, phi ⊢ _:(A)iota} {Gamma, phi ⊢ _:A} ∈ 𝕌{[i' j']})

Lemma: composition-type-lemma3
Gamma:j⊢. ∀phi:{Gamma ⊢ _:𝔽}. ∀A:{Gamma.𝕀 ⊢ _}. ∀u:{Gamma, phi.𝕀 ⊢ _:A}. ∀I,J:fset(ℕ). ∀f:J ⟶ I. ∀a:Gamma(I).
  ((u)<(s(f(a));<new-name(J)>)> iota
  ((u)<(s(a);<new-name(I)>)> iota)subset-trans(I+new-name(I);J+new-name(J);f,new-name(I)=new-name(J);s(phi(a)))
  ∈ {J+new-name(J),s(phi(f(a))) ⊢ _:(A)<(s(f(a));<new-name(J)>)> iota})

Lemma: composition-type-lemma4
Gamma:j⊢. ∀phi:{Gamma ⊢ _:𝔽}. ∀A:{Gamma.𝕀 ⊢ _}. ∀u:{Gamma, phi.𝕀 ⊢ _:A}. ∀I,J:fset(ℕ). ∀f:J ⟶ I. ∀rho:Gamma(I).
K:fset(ℕ). ∀g:J,phi(f(rho))(K).
  ((u)<(s(f(rho));<new-name(J)>)> iota((new-name(J)0) ⋅ g)
  ((u)<(s(rho);<new-name(I)>)> iota)subset-trans(I+new-name(I);J+new-name(J);f,new-name(I)=new-name(J);
                                                     s(phi(rho)))((new-name(J)0) ⋅ g)
  ∈ A(g((new-name(J)0)((s(f(rho));<new-name(J)>)))))

Lemma: composition-type-lemma5
[Gamma:j⊢]. ∀[phi:{Gamma ⊢ _:𝔽}]. ∀[A:{Gamma.𝕀 ⊢ _}]. ∀[u:{Gamma, phi.𝕀 ⊢ _:A}].
  ({Gamma ⊢ _:(A)[0(𝕀)][phi |⟶ (u)[0(𝕀)]]} ∈ 𝕌{[i j']})

Lemma: composition-type-lemma6
[Gamma:j⊢]. ∀[phi:{Gamma ⊢ _:𝔽}]. ∀[A:{Gamma.𝕀 ⊢ _}]. ∀[u:{Gamma, phi.𝕀 ⊢ _:A}].
  ({Gamma ⊢ _:(A)[1(𝕀)][phi |⟶ (u)[1(𝕀)]]} ∈ 𝕌{[i j']})

Definition: composition-term
comp cA [phi ⊢→ u] a0 ==
  λI,rho. (cA new-name(I) (s(rho);<new-name(I)>phi(rho) (u)<(s(rho);<new-name(I)>)> iota a0(rho))

Lemma: composition-term_wf
[Gamma:j⊢]. ∀[phi:{Gamma ⊢ _:𝔽}]. ∀[A:{Gamma.𝕀 ⊢ _}]. ∀[cA:Gamma.𝕀 ⊢ CompOp(A)]. ∀[u:{Gamma, phi.𝕀 ⊢ _:A}].
[a0:{Gamma ⊢ _:(A)[0(𝕀)][phi |⟶ (u)[0(𝕀)]]}].
  (comp cA [phi ⊢→ u] a0 ∈ {Gamma ⊢ _:(A)[1(𝕀)][phi |⟶ (u)[1(𝕀)]]})

Lemma: empty-context-eq-lemma
[Gamma:j⊢]. ∀[A,x,y:Top].  (x y ∈ {Gamma ⊢ _:A}) supposing ∀I:fset(ℕ). Gamma(I))

Lemma: empty-context-lemma
[Gamma:j⊢]. ∀[A,x:Top].  (x ∈ {Gamma ⊢ _:A}) supposing ∀I:fset(ℕ). Gamma(I))

Lemma: empty-context-subset-lemma1
[Gamma:j⊢]. ∀[A,x:Top].  (x ∈ {Gamma, 0(𝔽).𝕀 ⊢ _:A})

Lemma: empty-context-subset-lemma2
[Gamma:j⊢]. ∀[A,x:Top].  (x ∈ {Gamma, 0(𝔽) ⊢ _:A})

Lemma: empty-context-subset-lemma3
[Gamma:j⊢]. ∀[A,x,y:Top].  (x y ∈ {Gamma, 0(𝔽) ⊢ _:A})

Lemma: empty-context-subset-lemma3'
[Gamma:j⊢]. ∀[i:{Gamma ⊢ _:𝕀}]. ∀[A,x,y:Top].  (x y ∈ {Gamma, (i=0), (i=1) ⊢ _:A})

Lemma: empty-context-subset-lemma4
[Gamma:j⊢]. ∀[B:{Gamma ⊢ _}]. ∀[A,x,y:Top].  (x y ∈ {Gamma, 0(𝔽).B ⊢ _:A})

Lemma: empty-context-subset-lemma5
[Gamma:j⊢]. ∀[i:{Gamma ⊢ _:𝕀}]. ∀[x,y:Top × Top].  (x y ∈ {Gamma, (i=0), (i=1) ⊢ _})

Lemma: empty-context-subset-lemma6
[Gamma:j⊢]. ∀[x,y:Top × Top].  (x y ∈ {Gamma, 0(𝔽) ⊢ _})

Lemma: same-cubical-type-trivial_1
[X:j⊢]. ∀[i:{X ⊢ _:𝕀}]. ∀[x,y,B:Top].  X, (i=0), (i=1) ⊢ x=y:B

Lemma: case-type-1
[Gamma:j⊢]. ∀[phi:{Gamma ⊢ _:𝔽}].
  ∀[A:{Gamma ⊢ _}]. ∀[B:Top × Top].  Gamma ⊢ (if phi then else B) supposing phi 1(𝔽) ∈ {Gamma ⊢ _:𝔽}

Lemma: case-type-0
[Gamma:j⊢]. ∀[phi:{Gamma ⊢ _:𝔽}].
  ∀[A:Top × Top]. ∀[B:{Gamma ⊢ _}].  Gamma ⊢ (if phi then else B) supposing phi 0(𝔽) ∈ {Gamma ⊢ _:𝔽}

Lemma: case-term-0
[Gamma:j⊢]. ∀[phi:{Gamma ⊢ _:𝔽}]. ∀[A:{Gamma ⊢ _}]. ∀[u:Top]. ∀[v:{Gamma ⊢ _:A}].
  Gamma ⊢ (u ∨ v)=v:A supposing phi 0(𝔽) ∈ {Gamma ⊢ _:𝔽}

Lemma: case-term-0'
[Gamma:j⊢]. ∀[phi:{Gamma ⊢ _:𝔽}]. ∀[A:{Gamma ⊢ _}]. ∀[u:Top]. ∀[v,x:{Gamma ⊢ _:A}].
  (Gamma ⊢ (u ∨ v)=x:A) supposing ((x v ∈ {Gamma ⊢ _:A}) and (phi 0(𝔽) ∈ {Gamma ⊢ _:𝔽}))

Definition: case-endpoints
[r=0 ⊢→ a; r=1 ⊢→ b] ==  (a ∨ b)

Lemma: case-endpoints_wf
[G:j⊢]. ∀[A:{G ⊢ _}]. ∀[a,b:{G ⊢ _:A}]. ∀[r:{G ⊢ _:𝕀}].  ([r=0 ⊢→ a; r=1 ⊢→ b] ∈ {G, ((r=0) ∨ (r=1)) ⊢ _:A})

Lemma: case-endpoints-0
[G:j⊢]. ∀[A:{G ⊢ _}]. ∀[a:{G ⊢ _:A}]. ∀[b:Top].  ([0(𝕀)=0 ⊢→ a; 0(𝕀)=1 ⊢→ b] a ∈ {G ⊢ _:A})

Lemma: case-endpoints-1
[G:j⊢]. ∀[A:{G ⊢ _}]. ∀[a:Top]. ∀[b:{G ⊢ _:A}].  ([1(𝕀)=0 ⊢→ a; 1(𝕀)=1 ⊢→ b] b ∈ {G ⊢ _:A})

Lemma: csm-case-endpoints
[a,b,r,s:Top].  (([r=0 ⊢→ a; r=1 ⊢→ b])s [(r)s=0 ⊢→ (a)s; (r)s=1 ⊢→ (b)s])

Lemma: cubical-term-1-q1
[Gamma:j⊢]. ∀[B:{Gamma ⊢ _}]. ∀[z:{Gamma.𝕀 ⊢ _:(B)p}].  (((z)[1(𝕀)])p z ∈ {Gamma.𝕀(q=1) ⊢ _:(B)p})

Lemma: cubical-term-0-q0
[Gamma:j⊢]. ∀[B:{Gamma ⊢ _}]. ∀[z:{Gamma.𝕀 ⊢ _:(B)p}].  (((z)[0(𝕀)])p z ∈ {Gamma.𝕀(q=0) ⊢ _:(B)p})

Definition: rev-type-line
(A)- ==  (A)(p;1-(q))

Lemma: rev-type-line_wf
[Gamma:j⊢]. ∀[A:{Gamma.𝕀 ⊢ _}].  Gamma.𝕀 ⊢ (A)-

Lemma: csm-rev-type-line
[G,K:j⊢]. ∀[A:{G.𝕀 ⊢ _}]. ∀[tau:K j⟶ G].  (((A)-)tau+ ((A)tau+)- ∈ {K.𝕀 ⊢ _})

Lemma: rev-type-line-0
[Gamma:j⊢]. ∀[A:{Gamma.𝕀 ⊢ _}].  (((A)-)[0(𝕀)] (A)[1(𝕀)])

Lemma: rev-type-line-1
[Gamma:j⊢]. ∀[A:{Gamma.𝕀 ⊢ _}].  (((A)-)[1(𝕀)] (A)[0(𝕀)])

Lemma: rev-rev-type-line
[Gamma:j⊢]. ∀[A:{Gamma.𝕀 ⊢ _}].  (((A)-)- A ∈ {Gamma.𝕀 ⊢ _})

Definition: rev-type-line-comp
(cA)- ==  (cA)(p;1-(q))

Lemma: rev-type-line-comp_wf
[Gamma:j⊢]. ∀[A:{Gamma.𝕀 ⊢ _}]. ∀[cA:Gamma.𝕀 ⊢ CompOp(A)].  ((cA)- ∈ Gamma.𝕀 ⊢ CompOp((A)-))

Lemma: csm-rev-type-line-comp
[G,K:j⊢]. ∀[tau:K j⟶ G]. ∀[A:{G.𝕀 ⊢ _}]. ∀[cA:G.𝕀 ⊢ CompOp(A)].
  (((cA)-)tau+ ((cA)tau+)- ∈ K.𝕀 ⊢ CompOp(((A)-)tau+))

Definition: transport
transport(Gamma;a) ==  comp cA [0(𝔽) ⊢→ discr(⋅)] a

Lemma: transport_wf
[Gamma:j⊢]. ∀[A:{Gamma.𝕀 ⊢ _}]. ∀[cA:Gamma.𝕀 ⊢ CompOp(A)]. ∀[a:{Gamma ⊢ _:(A)[0(𝕀)]}].
  (transport(Gamma;a) ∈ {Gamma ⊢ _:(A)[1(𝕀)]})

Definition: transport-fun
transport-fun(Gamma;A;cA) ==  transport(Gamma.(A)[0(𝕀)];q))

Lemma: transport-fun_wf
[Gamma:j⊢]. ∀[A:{Gamma.𝕀 ⊢ _}]. ∀[cA:Gamma.𝕀 ⊢ CompOp(A)].
  (transport-fun(Gamma;A;cA) ∈ {Gamma ⊢ _:((A)[0(𝕀)] ⟶ (A)[1(𝕀)])})

Definition: transport-const
transport-const(G;cA;a) ==  transport(G;a)

Lemma: transport-const_wf
[G:j⊢]. ∀[A:{G ⊢ _}]. ∀[cA:G ⊢ CompOp(A)]. ∀[a:{G ⊢ _:A}].  (transport-const(G;cA;a) ∈ {G ⊢ _:A})

Definition: const-transport-fun
ConstTrans(A) ==  cubical-lam(G;transport-const(G.A;(cA)p;q))

Lemma: const-transport-fun_wf
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[cA:Gamma ⊢ CompOp(A)].  (ConstTrans(A) ∈ {Gamma ⊢ _:(A ⟶ A)})

Definition: rev-transport-fun
rev-transport-fun(Gamma;A;cA) ==  transport-fun(Gamma;(A)-;(cA)(p;1-(q)))

Lemma: rev-transport-fun_wf
[Gamma:j⊢]. ∀[A:{Gamma.𝕀 ⊢ _}]. ∀[cA:Gamma.𝕀 ⊢ CompOp(A)].
  (rev-transport-fun(Gamma;A;cA) ∈ {Gamma ⊢ _:((A)[1(𝕀)] ⟶ (A)[0(𝕀)])})

Definition: cubical-contr
cubical-contr(Gamma; A; cA; p; phi; u) ==  comp (cA)p [phi ⊢→ (app(p.2; u))p q] p.1

Lemma: cubical-contr_wf
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[cA:Gamma ⊢ CompOp(A)]. ∀[p:{Gamma ⊢ _:Contractible(A)}]. ∀[phi:{Gamma ⊢ _:𝔽}].
[u:{Gamma, phi ⊢ _:(A)iota}].
  (cubical-contr(Gamma; A; cA; p; phi; u) ∈ {Gamma ⊢ _:A[phi |⟶ u]})

Lemma: csm-subtype-cubical-subset
[Gamma:j⊢]. ∀[I:fset(ℕ)]. ∀[psi:𝔽(I)].  (formal-cube(I) j⟶ Gamma ⊆I,psi j⟶ Gamma)

Lemma: cubical-type-subtype-cubical-subset
[I:fset(ℕ)]. ∀[psi:𝔽(I)].  ({formal-cube(I) ⊢ _} ⊆{I,psi ⊢ _})

Lemma: cubical-term-subtype-cubical-subset
[I:fset(ℕ)]. ∀[psi:𝔽(I)]. ∀[T:{formal-cube(I) ⊢ _}].  ({formal-cube(I) ⊢ _:T} ⊆{I,psi ⊢ _:T})

Lemma: context-map_wf_cubical-subset
[Gamma:j⊢]. ∀[I:fset(ℕ)]. ∀[rho:Gamma(I)]. ∀[psi:𝔽(I)].  (<rho> ∈ I,psi j⟶ Gamma)

Definition: composition-function
composition-function{j:l,i:l}(Gamma;A) ==
  H:CubicalSet{j}
  ⟶ sigma:H.𝕀 j⟶ Gamma
  ⟶ phi:{H ⊢ _:𝔽}
  ⟶ u:{H, phi.𝕀 ⊢ _:(A)sigma}
  ⟶ a0:{H ⊢ _:((A)sigma)[0(𝕀)][phi |⟶ (u)[0(𝕀)]]}
  ⟶ {H ⊢ _:((A)sigma)[1(𝕀)][phi |⟶ (u)[1(𝕀)]]}

Lemma: composition-function_wf
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}].  (composition-function{j:l,i:l}(Gamma;A) ∈ 𝕌{[i' j'']})

Lemma: composition-function-cumulativity
Gamma:j⊢. ∀Z:{Gamma ⊢ _}.  (composition-function{[i j]:l, i:l}(Gamma; Z) ⊆composition-function{j:l,i:l}(Gamma;Z))

Definition: uniform-comp-function
uniform-comp-function{j:l, i:l}(Gamma; A; comp) ==
  ∀H,K:j⊢. ∀tau:K j⟶ H. ∀sigma:H.𝕀 j⟶ Gamma. ∀phi:{H ⊢ _:𝔽}. ∀u:{H, phi.𝕀 ⊢ _:(A)sigma}.
  ∀a0:{H ⊢ _:((A)sigma)[0(𝕀)][phi |⟶ (u)[0(𝕀)]]}.
    ((comp sigma phi a0)tau
    (comp sigma tau+ (phi)tau (u)tau+ (a0)tau)
    ∈ {K ⊢ _:(((A)sigma)[1(𝕀)])tau[(phi)tau |⟶ ((u)[1(𝕀)])tau]})

Lemma: uniform-comp-function_wf
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[comp:composition-function{j:l,i:l}(Gamma;A)].
  (uniform-comp-function{j:l, i:l}(Gamma; A; comp) ∈ ℙ{[i' j'']})

Lemma: uniform-comp-function-cumulativity
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[comp:composition-function{[i j]:l, i:l}(Gamma; A)].
  (uniform-comp-function{[i j]:l, i:l}(Gamma; A; comp)  uniform-comp-function{j:l, i:l}(Gamma; A; comp))

Definition: composition-structure
Gamma ⊢ Compositon(A) ==
  {comp:composition-function{j:l,i:l}(Gamma;A)| uniform-comp-function{j:l, i:l}(Gamma; A; comp)} 

Lemma: composition-structure_wf
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}].  (Gamma ⊢ Compositon(A) ∈ 𝕌{[i' j'']})

Lemma: composition-structure-cumulativity
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}].  (Gamma +⊢ Compositon(A) ⊆Gamma ⊢ Compositon(A))

Lemma: composition-structure-equal
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[c1,c2:Gamma ⊢ Compositon(A)].
  c1 c2 ∈ Gamma ⊢ Compositon(A) 
  supposing ∀H:j⊢. ∀sigma:H.𝕀 j⟶ Gamma. ∀phi:{H ⊢ _:𝔽}. ∀u:{H, phi.𝕀 ⊢ _:(A)sigma}.
            ∀a0:{H ⊢ _:((A)sigma)[0(𝕀)][phi |⟶ (u)[0(𝕀)]]}.
              ((c1 sigma phi a0) (c2 sigma phi a0) ∈ {H ⊢ _:((A)sigma)[1(𝕀)]})

Lemma: composition-function-subset
[Y,X:j⊢].
  ∀[B:{X ⊢ _}]. (composition-function{j:l,i:l}(X;B) ⊆composition-function{j:l,i:l}(Y;B)) 
  supposing sub_cubical_set{j:l}(Y; X)

Lemma: composition-structure-subset
[Y,X:j⊢].  ∀[B:{X ⊢ _}]. (X ⊢ Compositon(B) ⊆Y ⊢ Compositon(B)) supposing sub_cubical_set{j:l}(Y; X)

Lemma: composition-in-subset
[G:j⊢]. ∀[A:{G ⊢ _}]. ∀[cA:G ⊢ Compositon(A)]. ∀[H1,H2:j⊢].
  ∀[sigma:H1.𝕀 j⟶ G]. ∀[phi:{H1 ⊢ _:𝔽}]. ∀[u:{H1, phi.𝕀 ⊢ _:(A)sigma}].
  ∀[a0:{H1 ⊢ _:((A)sigma)[0(𝕀)][phi |⟶ (u)[0(𝕀)]]}].
    ((cA H1 sigma phi a0) (cA H2 sigma phi a0) ∈ {H2 ⊢ _:((A)sigma)[1(𝕀)]}) 
  supposing sub_cubical_set{j:l}(H2; H1)

Definition: comp-op-to-comp-fun
cop-to-cfun(cA) ==  ⌜λH,sigma,phi,u,a0. comp (cA)sigma [phi ⊢→ u] a0⌝

Lemma: comp-op-to-comp-fun_wf
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[cA:Gamma ⊢ CompOp(A)].  (cop-to-cfun(cA) ∈ composition-function{j:l,i:l}(Gamma;A))

Lemma: face-type-comp-at-lemma
[H:j⊢]. ∀[phi:{H ⊢ _:𝔽}]. ∀[I,J:fset(ℕ)]. ∀[i:ℕ]. ∀[f:J ⟶ I+i]. ∀[v:H(I)].  (phi(f(s(v))) f(s(phi(v))) ∈ 𝔽(f))

Lemma: composition-term-uniformity
[H,K:j⊢]. ∀[tau:K j⟶ H]. ∀[phi:{H ⊢ _:𝔽}]. ∀[A:{H.𝕀 ⊢ _}]. ∀[u:{H, phi.𝕀 ⊢ _:A}].
[a0:{H ⊢ _:(A)[0(𝕀)][phi |⟶ (u)[0(𝕀)]]}]. ∀[cA:H.𝕀 ⊢ CompOp(A)].
  ((comp cA [phi ⊢→ u] a0)tau comp (cA)tau+ [(phi)tau ⊢→ (u)tau+] (a0)tau ∈ {K ⊢ _:((A)[1(𝕀)])tau})

Lemma: csm-transport
[Gamma:j⊢]. ∀[A:{Gamma.𝕀 ⊢ _}]. ∀[cA:Gamma.𝕀 ⊢ CompOp(A)]. ∀[a:{Gamma ⊢ _:(A)[0(𝕀)]}]. ∀[H:j⊢]. ∀[s:H j⟶ Gamma].
  ((transport(Gamma;a))s transport(H;(a)s) ∈ {H ⊢ _:((A)[1(𝕀)])s})

Lemma: csm-transport-fun
[Gamma:j⊢]. ∀[A:{Gamma.𝕀 ⊢ _}]. ∀[cA:Gamma.𝕀 ⊢ CompOp(A)]. ∀[H:j⊢]. ∀[s:H j⟶ Gamma].
  ((transport-fun(Gamma;A;cA))s transport-fun(H;(A)s+;(cA)s+) ∈ {H ⊢ _:(((A)s+)[0(𝕀)] ⟶ ((A)s+)[1(𝕀)])})

Lemma: csm-comp-op-to-comp-fun
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[cA:Gamma ⊢ CompOp(A)]. ∀[H,K:j⊢]. ∀[tau:K j⟶ H]. ∀[sigma:H.𝕀 j⟶ Gamma].
[phi:{H ⊢ _:𝔽}]. ∀[u:{H, phi.𝕀 ⊢ _:(A)sigma}]. ∀[a0:{H ⊢ _:((A)sigma)[0(𝕀)][phi |⟶ (u)[0(𝕀)]]}].
  ((cop-to-cfun(cA) sigma phi a0)tau
  (cop-to-cfun(cA) sigma tau+ (phi)tau (u)tau+ (a0)tau)
  ∈ {K ⊢ _:(((A)sigma)[1(𝕀)])tau[(phi)tau |⟶ ((u)[1(𝕀)])tau]})

Lemma: comp-op-to-comp-fun_wf2
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[cA:Gamma ⊢ CompOp(A)].  (cop-to-cfun(cA) ∈ Gamma ⊢ Compositon(A))

Lemma: composition-op-implies-composition-structure
Gamma:j⊢. ∀A:{Gamma ⊢ _}. ∀cA:Gamma ⊢ CompOp(A).  Gamma ⊢ Compositon(A)

Lemma: csm-ap-term-cube+
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[I:fset(ℕ)]. ∀[i:{i:ℕ| ¬i ∈ I} ]. ∀[rho:Gamma(I+i)]. ∀[phi:𝔽(I)].
[u:{I+i,s(phi) ⊢ _:(A)<rho> iota}].
  ((u)cube+(I;i) ∈ {formal-cube(I), canonical-section(();𝔽;I;⋅;phi).𝕀 ⊢ _:(A)<rho> cube+(I;i)})

Lemma: canonical-section-cubical-path-0
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[I:fset(ℕ)]. ∀[i:{i:ℕ| ¬i ∈ I} ]. ∀[rho:Gamma(I+i)]. ∀[phi:𝔽(I)].
[u:{I+i,s(phi) ⊢ _:(A)<rho> iota}]. ∀[a0:cubical-path-0(Gamma;A;I;i;rho;phi;u)].
  (canonical-section(Gamma;A;I;(i0)(rho);a0)
   ∈ {formal-cube(I) ⊢ _:((A)<rho> cube+(I;i))[0(𝕀)][canonical-section(();𝔽;I;⋅;phi) |⟶ ((u)cube+(I;i))[0(𝕀)]]})

Lemma: constrained-cubical-term-to-cubical-path-1
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[I:fset(ℕ)]. ∀[i:{i:ℕ| ¬i ∈ I} ]. ∀[rho:Gamma(I+i)].
  ∀phi:𝔽(I)
    ∀[u:{I+i,s(phi) ⊢ _:(A)<rho> iota}].
    ∀[v:{formal-cube(I) ⊢ _:((A)<rho> cube+(I;i))[1(𝕀)][canonical-section(();𝔽;I;⋅;phi) |⟶ ((u)cube+(I;i))[1(𝕀)]]}].
      (v(1) ∈ cubical-path-1(Gamma;A;I;i;rho;phi;u))

Definition: comp-fun-to-comp-op1
comp-fun-to-comp-op1(Gamma;A;comp) ==
  λI,i,rho,phi,u,a0. (comp formal-cube(I) <rho> cube+(I;i) canonical-section(();𝔽;I;⋅;phi) (u)cube+(I;i) 
                      canonical-section(Gamma;A;I;(i0)(rho);a0))

Lemma: comp-fun-to-comp-op1_wf
Gamma:j⊢. ∀A:{Gamma ⊢ _}.
  ∀[comp:composition-function{j:l,i:l}(Gamma;A)]
    (comp-fun-to-comp-op1(Gamma;A;comp) ∈ I:fset(ℕ)
     ⟶ i:{i:ℕ| ¬i ∈ I} 
     ⟶ rho:Gamma(I+i)
     ⟶ phi:𝔽(I)
     ⟶ u:{I+i,s(phi) ⊢ _:(A)<rho> iota}
     ⟶ cubical-path-0(Gamma;A;I;i;rho;phi;u)
     ⟶ {formal-cube(I) ⊢ _:((A)<rho> cube+(I;i))[1(𝕀)][canonical-section(();𝔽;I;⋅;phi) |⟶ ((u)cube+(I;i))[1(𝕀)]]})

Definition: comp-fun-to-comp-op
cfun-to-cop(Gamma;A;comp) ==  λI,i,rho,phi,u,a0. comp-fun-to-comp-op1(Gamma;A;comp) rho phi a0(1)

Lemma: comp-fun-to-comp-op_wf1
Gamma:j⊢. ∀A:{Gamma ⊢ _}.
  ∀[comp:composition-function{j:l,i:l}(Gamma;A)]
    (cfun-to-cop(Gamma;A;comp) ∈ I:fset(ℕ)
     ⟶ i:{i:ℕ| ¬i ∈ I} 
     ⟶ rho:Gamma(I+i)
     ⟶ phi:𝔽(I)
     ⟶ u:{I+i,s(phi) ⊢ _:(A)<rho> iota}
     ⟶ cubical-path-0(Gamma;A;I;i;rho;phi;u)
     ⟶ cubical-path-1(Gamma;A;I;i;rho;phi;u))

Lemma: comp-fun-to-comp-op_wf
Gamma:j⊢. ∀A:{Gamma ⊢ _}.  ∀[comp:Gamma ⊢ Compositon(A)]. (cfun-to-cop(Gamma;A;comp) ∈ Gamma ⊢ CompOp(A))

Lemma: composition-structure-implies-composition-op
Gamma:j⊢. ∀A:{Gamma ⊢ _}.  (Gamma ⊢ Compositon(A)  Gamma ⊢ CompOp(A))

Lemma: comp-op-to-comp-fun-inverse
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[cA:Gamma ⊢ CompOp(A)].  (cfun-to-cop(Gamma;A;cop-to-cfun(cA)) cA ∈ Gamma ⊢ CompOp(A))

Lemma: comp-fun-to-comp-op-inverse
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[cA:Gamma ⊢ Compositon(A)].
  (cop-to-cfun(cfun-to-cop(Gamma;A;cA)) cA ∈ Gamma ⊢ Compositon(A))

Definition: csm-comp-structure
(cA)tau ==  λH,sigma,phi,u,a0. (cA tau sigma phi a0)

Lemma: csm-comp-structure_wf
[Gamma,Delta:j⊢]. ∀[tau:Delta j⟶ Gamma]. ∀[A:{Gamma ⊢ _}]. ∀[cA:Gamma ⊢ Compositon(A)].
  ((cA)tau ∈ Delta ⊢ Compositon((A)tau))

Lemma: csm-comp-structure_wf2
[Gamma,Delta:j⊢]. ∀[tau:Delta j⟶ Gamma]. ∀[A:{Gamma ⊢ _}]. ∀[cA:Gamma +⊢ Compositon(A)].
  ((cA)tau ∈ Delta +⊢ Compositon((A)tau))

Lemma: csm-comp-structure-composition-function
[Gamma,Delta:j⊢]. ∀[tau:Delta j⟶ Gamma]. ∀[A:{Gamma ⊢ _}]. ∀[cA:Gamma ⊢ Compositon(A)].
  ((cA)tau ∈ composition-function{j:l,i:l}(Delta;(A)tau))

Lemma: csm-comp-structure-subset
[Gamma,Delta,tau,cA,phi:Top].  ((cA)tau (cA)tau)

Lemma: csm-comp-structure-id
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[cA:Gamma ⊢ Compositon(A)]. ∀[tau:Gamma j⟶ Gamma].
  (cA)tau cA ∈ Gamma ⊢ Compositon(A) supposing tau 1(Gamma) ∈ Gamma j⟶ Gamma

Lemma: csm-comp-op-to-comp-fun-sq
[Gamma,Delta,tau,cA:Top].  ((cop-to-cfun(cA))tau cop-to-cfun((cA)tau))

Lemma: csm-comp-fun-to-comp-op
[Gamma,K:j⊢]. ∀[tau:K j⟶ Gamma]. ∀[A:{Gamma ⊢ _}]. ∀[cA:Gamma ⊢ Compositon(A)].
  ((cfun-to-cop(Gamma;A;cA))tau cfun-to-cop(K;(A)tau;(cA)tau) ∈ K ⊢ CompOp((A)tau))

Lemma: subset-comp-structure
[X,Y:j⊢]. ∀[T:{Y ⊢ _}].  Y ⊢ Compositon(T) ⊆X ⊢ Compositon(T) supposing sub_cubical_set{j:l}(X; Y)

Definition: rev-type-comp
rev-type-comp(Gamma;cA) ==  (cA)(p;1-(q))

Lemma: rev-type-comp_wf
[Gamma:j⊢]. ∀[A:{Gamma.𝕀 ⊢ _}]. ∀[cA:Gamma.𝕀 ⊢ Compositon(A)].  (rev-type-comp(Gamma;cA) ∈ Gamma.𝕀 ⊢ Compositon((A)-))

Definition: comp_term
comp cA [phi ⊢→ u] a0 ==  cA Gamma 1(Gamma.𝕀phi a0

Lemma: comp_term_wf
[Gamma:j⊢]. ∀[phi:{Gamma ⊢ _:𝔽}]. ∀[A:{Gamma.𝕀 ⊢ _}]. ∀[cA:composition-function{j:l,i:l}(Gamma.𝕀;A)].
[u:{Gamma, phi.𝕀 ⊢ _:A}]. ∀[a0:{Gamma ⊢ _:(A)[0(𝕀)][phi |⟶ (u)[0(𝕀)]]}].
  (comp cA [phi ⊢→ u] a0 ∈ {Gamma ⊢ _:(A)[1(𝕀)][phi |⟶ (u)[1(𝕀)]]})

Lemma: comp_term-composition-term
[Gamma,phi,cA,u,a0:Top].  (comp cop-to-cfun(cA) [phi ⊢→ u] a0 comp cA [phi ⊢→ u] a0)

Lemma: csm-comp_term
[Gamma:j⊢]. ∀[phi:{Gamma ⊢ _:𝔽}]. ∀[A:{Gamma.𝕀 ⊢ _}]. ∀[cA:Gamma.𝕀 ⊢ Compositon(A)]. ∀[u:{Gamma, phi.𝕀 ⊢ _:A}].
[a0:{Gamma ⊢ _:(A)[0(𝕀)][phi |⟶ (u)[0(𝕀)]]}]. ∀[Delta:j⊢]. ∀[s:Delta j⟶ Gamma].
  ((comp cA [phi ⊢→ u] a0)s comp (cA)s+ [(phi)s ⊢→ (u)s+] (a0)s ∈ {Delta ⊢ _:((A)s+)[1(𝕀)][(phi)s |⟶ ((u)s+)[1(𝕀)]]})

Definition: transprt
transprt(G;cA;a0) ==  comp cA [0(𝔽) ⊢→ discr(⋅)] a0

Lemma: transprt_wf
[Gamma:j⊢]. ∀[A:{Gamma.𝕀 ⊢ _}]. ∀[cA:Gamma.𝕀 ⊢ Compositon(A)]. ∀[a:{Gamma ⊢ _:(A)[0(𝕀)]}].
  (transprt(Gamma;cA;a) ∈ {Gamma ⊢ _:(A)[1(𝕀)]})

Lemma: equals-transprt
[Gamma:j⊢]. ∀[A:{Gamma.𝕀 ⊢ _}]. ∀[cA:Gamma.𝕀 +⊢ Compositon(A)]. ∀[a:{Gamma ⊢ _:(A)[0(𝕀)]}]. ∀[xx:Top].
  (comp cA [0(𝔽) ⊢→ xx] transprt(Gamma;cA;a) ∈ {Gamma ⊢ _:(A)[1(𝕀)]})

Definition: transprt-fun
transprt-fun(Gamma;A;cA) ==  transprt(Gamma.(A)[0(𝕀)];(cA)p+;q))

Lemma: transprt-fun_wf
[Gamma:j⊢]. ∀[A:{Gamma.𝕀 ⊢ _}]. ∀[cA:Gamma.𝕀 +⊢ Compositon(A)].
  (transprt-fun(Gamma;A;cA) ∈ {Gamma ⊢ _:((A)[0(𝕀)] ⟶ (A)[1(𝕀)])})

Lemma: csm-transprt
[Gamma:j⊢]. ∀[A:{Gamma.𝕀 ⊢ _}]. ∀[cA:Gamma.𝕀 ⊢ Compositon(A)]. ∀[a:{Gamma ⊢ _:(A)[0(𝕀)]}]. ∀[H:j⊢]. ∀[s:H j⟶ Gamma].
  ((transprt(Gamma;cA;a))s transprt(H;(cA)s+;(a)s) ∈ {H ⊢ _:((A)[1(𝕀)])s})

Lemma: csm-transprt-fun
[Gamma:j⊢]. ∀[A:{Gamma.𝕀 ⊢ _}]. ∀[cA:Gamma.𝕀 +⊢ Compositon(A)]. ∀[H:j⊢]. ∀[s:H j⟶ Gamma].
  ((transprt-fun(Gamma;A;cA))s transprt-fun(H;(A)s+;(cA)s+) ∈ {H ⊢ _:(((A)s+)[0(𝕀)] ⟶ ((A)s+)[1(𝕀)])})

Definition: transprt-const
transprt-const(G;cA;a) ==  transprt(G;(cA)p;a)

Lemma: transprt-const_wf
[G:j⊢]. ∀[A:{G ⊢ _}]. ∀[cA:G +⊢ Compositon(A)]. ∀[a:{G ⊢ _:A}].  (transprt-const(G;cA;a) ∈ {G ⊢ _:A})

Lemma: csm-transprt-const
[G:j⊢]. ∀[A:{G ⊢ _}]. ∀[cA:G +⊢ Compositon(A)]. ∀[a:{G ⊢ _:A}]. ∀[H:j⊢]. ∀[s:H j⟶ G].
  ((transprt-const(G;cA;a))s transprt-const(H;(cA)s;(a)s) ∈ {H ⊢ _:(A)s})

Definition: comp_trm
comp_trm(Gamma;cA;phi;u;a0) ==  comp cA [phi ⊢→ u] a0

Lemma: comp_trm_wf
[Gamma:j⊢]. ∀[phi:{Gamma ⊢ _:𝔽}]. ∀[A:{Gamma.𝕀 ⊢ _}]. ∀[cA:composition-function{j:l,i:l}(Gamma.𝕀;A)].
[u:{Gamma, phi.𝕀 ⊢ _:A}]. ∀[a0:{Gamma ⊢ _:(A)[0(𝕀)][phi |⟶ (u)[0(𝕀)]]}].
  (comp_trm(Gamma;
            cA;
            phi;
            u;
            a0) ∈ {Gamma ⊢ _:(A)[1(𝕀)]})

Lemma: comp_term-subset
[Gamma:j⊢]. ∀[phi:{Gamma ⊢ _:𝔽}]. ∀[A:{Gamma.𝕀 ⊢ _}]. ∀[cA:Gamma.𝕀 ⊢ Compositon(A)]. ∀[u:{Gamma, phi.𝕀 ⊢ _:A}].
[a0:{Gamma ⊢ _:(A)[0(𝕀)][phi |⟶ (u)[0(𝕀)]]}]. ∀[psi:{Gamma ⊢ _:𝔽}].
  (comp cA [phi ⊢→ u] a0 comp cA [phi ⊢→ u] a0 ∈ {Gamma, psi ⊢ _:(A)[1(𝕀)]})

Definition: filling-function
filling-function{j:l, i:l}(Gamma;A) ==
  H:CubicalSet{j}
  ⟶ sigma:H.𝕀 j⟶ Gamma
  ⟶ phi:{H ⊢ _:𝔽}
  ⟶ u:{H.𝕀(phi)p ⊢ _:(A)sigma}
  ⟶ a0:{H ⊢ _:((A)sigma)[0(𝕀)][phi |⟶ u[0]]}
  ⟶ {H.𝕀 ⊢ _:(A)sigma[(phi)p |⟶ u]}

Lemma: filling-function_wf
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}].  (filling-function{j:l, i:l}(Gamma;A) ∈ 𝕌{[i' j'']})

Definition: uniform-filling-function
uniform-filling-function{j:l, i:l}(Gamma;A;fill) ==
  ∀H,K:j⊢. ∀tau:K j⟶ H. ∀sigma:H.𝕀 j⟶ Gamma. ∀phi:{H ⊢ _:𝔽}. ∀u:{H.𝕀(phi)p ⊢ _:(A)sigma}.
  ∀a0:{H ⊢ _:((A)sigma)[0(𝕀)][phi |⟶ u[0]]}.
    ((fill sigma phi a0)tau+ (fill sigma tau+ (phi)tau (u)tau+ (a0)tau) ∈ {K.𝕀 ⊢ _:((A)sigma)tau+})

Lemma: uniform-filling-function_wf
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[comp:filling-function{j:l, i:l}(Gamma;A)].
  (uniform-filling-function{j:l, i:l}(Gamma;A;comp) ∈ ℙ{[i' j'']})

Definition: filling-structure
Gamma ⊢ Filling(A) ==  {fill:filling-function{j:l, i:l}(Gamma;A)| uniform-filling-function{j:l, i:l}(Gamma;A;fill)} 

Lemma: filling-structure_wf
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}].  (Gamma ⊢ Filling(A) ∈ 𝕌{[i' j'']})

Definition: csm-m
==  λI,c. let a,s in let rho,r in (rho;r ∧ s)

Lemma: csm-m_wf
[Gamma:j⊢]. (m ∈ Gamma.𝕀.𝕀 j⟶ Gamma.𝕀)

Lemma: cc-fst-comp-csm-m-term
[H:j⊢]. ∀[phi:{H ⊢ _:𝔽}].  (((phi)p)m ((phi)p)p ∈ {H.𝕀.𝕀 ⊢ _:𝔽})

Lemma: 0-comp-cc-fst-comp-m
[H:j⊢]. ([0(𝕀)] m ∈ H.𝕀.𝕀((q=0))p j⟶ H.𝕀)

Lemma: csm-m-comp-0
[H:j⊢]. ([0(𝕀)] [0(𝕀)] ∈ H.𝕀 ij⟶ H.𝕀)

Lemma: csm-m-comp-1
[H:j⊢]. (m [1(𝕀)] 1(H.𝕀) ∈ H.𝕀 ij⟶ H.𝕀)

Lemma: csm+-comp-m
[H,K:j⊢]. ∀[tau:K j⟶ H].  (m tau++ tau+ m ∈ K.𝕀.𝕀 ij⟶ H.𝕀)

Lemma: cc-fst+-comp-0
[G:j⊢]. (p+ [0(𝕀)] [0(𝕀)] p ∈ G.𝕀 ij⟶ G.𝕀)

Lemma: cc-fst+-0-type
[G:j⊢]. ∀[A:{G.𝕀 ⊢ _}].  (((A)[0(𝕀)])p ((A)p+)[0(𝕀)] ∈ {G.𝕀 ⊢ _})

Lemma: cc-fst+-comp-1
[G:j⊢]. (p+ [1(𝕀)] [1(𝕀)] p ∈ G.𝕀 ij⟶ G.𝕀)

Lemma: cc-fst+-1-type
[G:j⊢]. ∀[A:{G.𝕀 ⊢ _}].  (((A)[1(𝕀)])p ((A)p+)[1(𝕀)] ∈ {G.𝕀 ⊢ _})

Definition: comp-to-fill
comp-to-fill(Gamma;cA) ==  λH,sigma,phi,u,a0. (cA H.𝕀 sigma ((phi)p ∨ (q=0)) ((u)m ∨ ((a0)p)m) (a0)p)

Lemma: comp-to-fill_wf
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[cA:composition-function{j:l,i:l}(Gamma;A)].
  (comp-to-fill(Gamma;cA) ∈ filling-function{j:l, i:l}(Gamma;A))

Lemma: comp-to-fill_wf2
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[cA:Gamma ⊢ Compositon(A)].  (comp-to-fill(Gamma;cA) ∈ Gamma ⊢ Filling(A))

Lemma: composition-implies-filling-structure
Gamma:j⊢. ∀A:{Gamma ⊢ _}.  (Gamma ⊢ Compositon(A)  Gamma ⊢ Filling(A))

Definition: fill_term
fill cA [phi ⊢→ u] a0 ==  comp-to-fill(Gamma.𝕀;cA) Gamma 1(Gamma.𝕀phi a0

Lemma: fill_term_wf
[Gamma:j⊢]. ∀[phi:{Gamma ⊢ _:𝔽}]. ∀[A:{Gamma.𝕀 ⊢ _}]. ∀[cA:composition-function{j:l,i:l}(Gamma.𝕀;A)].
[u:{Gamma.𝕀(phi)p ⊢ _:A}]. ∀[a0:{Gamma ⊢ _:(A)[0(𝕀)][phi |⟶ u[0]]}].
  (fill cA [phi ⊢→ u] a0 ∈ {Gamma.𝕀 ⊢ _:A[(phi)p |⟶ u]})

Lemma: csm-fill_term
[Gamma:j⊢]. ∀[phi:{Gamma ⊢ _:𝔽}]. ∀[A:{Gamma.𝕀 ⊢ _}]. ∀[cA:Gamma.𝕀 ⊢ Compositon(A)]. ∀[u:{Gamma.𝕀(phi)p ⊢ _:A}].
[a0:{Gamma ⊢ _:(A)[0(𝕀)][phi |⟶ u[0]]}]. ∀[Delta:j⊢]. ∀[s:Delta j⟶ Gamma].
  ((fill cA [phi ⊢→ u] a0)s+ fill (cA)s+ [(phi)s ⊢→ (u)s+] (a0)s ∈ {Delta.𝕀 ⊢ _:(A)s+[((phi)s)p |⟶ (u)s+]})

Lemma: fill_term_1
[H:j⊢]. ∀[phi:{H ⊢ _:𝔽}]. ∀[T:{H.𝕀 ⊢ _}]. ∀[u:{H.𝕀(phi)p ⊢ _:T}]. ∀[a0:{H ⊢ _:(T)[0(𝕀)][phi |⟶ u[0]]}].
[cT:H.𝕀 ⊢ Compositon(T)].
  ((fill cT [phi ⊢→ u] a0)[1(𝕀)] comp cT [phi ⊢→ u] a0 ∈ {H ⊢ _:(T)[1(𝕀)]})

Lemma: fill_term_0
[H:j⊢]. ∀[phi:{H ⊢ _:𝔽}]. ∀[T:{H.𝕀 ⊢ _}]. ∀[u:{H.𝕀(phi)p ⊢ _:T}]. ∀[a0:{H ⊢ _:(T)[0(𝕀)][phi |⟶ u[0]]}].
[cT:H.𝕀 ⊢ Compositon(T)].
  ((fill cT [phi ⊢→ u] a0)[0(𝕀)] a0 ∈ {H ⊢ _:(T)[0(𝕀)]})

Definition: rev_fill_term
rev_fill_term(Gamma;cA;phi;u;a1) ==  (fill rev-type-comp(Gamma;cA) [phi ⊢→ (u)(p;1-(q))] a1)(p;1-(q))

Lemma: rev_fill_term_wf
[Gamma:j⊢]. ∀[phi:{Gamma ⊢ _:𝔽}]. ∀[A:{Gamma.𝕀 ⊢ _}]. ∀[cA:Gamma.𝕀 ⊢ Compositon(A)]. ∀[u:{Gamma.𝕀(phi)p ⊢ _:A}].
[a1:{Gamma ⊢ _:(A)[1(𝕀)][phi |⟶ u[1]]}].
  (rev_fill_term(Gamma;cA;phi;u;a1) ∈ {Gamma.𝕀 ⊢ _:A[(phi)p |⟶ u]})

Lemma: rev_fill_term_1
[Gamma:j⊢]. ∀[phi:{Gamma ⊢ _:𝔽}]. ∀[A:{Gamma.𝕀 ⊢ _}]. ∀[cA:Gamma.𝕀 ⊢ Compositon(A)]. ∀[u:{Gamma.𝕀(phi)p ⊢ _:A}].
[a1:{Gamma ⊢ _:(A)[1(𝕀)][phi |⟶ u[1]]}].
  ((rev_fill_term(Gamma;cA;phi;u;a1))[1(𝕀)] a1 ∈ {Gamma ⊢ _:(A)[1(𝕀)]})

Lemma: rev_fill_term_0
[Gamma:j⊢]. ∀[phi:{Gamma ⊢ _:𝔽}]. ∀[A:{Gamma.𝕀 ⊢ _}]. ∀[cA:Gamma.𝕀 ⊢ Compositon(A)]. ∀[u:{Gamma.𝕀(phi)p ⊢ _:A}].
[a1:{Gamma ⊢ _:(A)[1(𝕀)][phi |⟶ u[1]]}].
  ((rev_fill_term(Gamma;cA;phi;u;a1))[0(𝕀)]
  comp rev-type-comp(Gamma;cA) [phi ⊢→ (u)(p;1-(q))] a1
  ∈ {Gamma ⊢ _:(A)[0(𝕀)]})

Lemma: csm-rev_fill_term
[Gamma:j⊢]. ∀[phi:{Gamma ⊢ _:𝔽}]. ∀[A:{Gamma.𝕀 ⊢ _}]. ∀[cA:Gamma.𝕀 ⊢ Compositon(A)]. ∀[u:{Gamma.𝕀(phi)p ⊢ _:A}].
[a1:{Gamma ⊢ _:(A)[1(𝕀)][phi |⟶ u[1]]}]. ∀[Delta:j⊢]. ∀[s:Delta j⟶ Gamma].
  ((rev_fill_term(Gamma;cA;phi;u;a1))s+
  rev_fill_term(Delta;(cA)s+;(phi)s;(u)s+;(a1)s)
  ∈ {Delta.𝕀 ⊢ _:(A)s+[((phi)s)p |⟶ (u)s+]})

Lemma: trivial-constrained-term
[Gamma:j⊢]. ∀[B:{Gamma ⊢ _}]. ∀[xx:Top].  ({Gamma ⊢ _:B} ⊆{Gamma ⊢ _:B[0(𝔽|⟶ xx]})

Definition: revfill
revfill(Gamma;cA;a1) ==  rev_fill_term(Gamma;cA;0(𝔽);discr(⋅);a1)

Lemma: revfill_wf
[Gamma:j⊢]. ∀[A:{Gamma.𝕀 ⊢ _}]. ∀[cA:Gamma.𝕀 ⊢ Compositon(A)]. ∀[a1:{Gamma ⊢ _:(A)[1(𝕀)]}].
  (revfill(Gamma;cA;a1) ∈ {Gamma.𝕀 ⊢ _:A})

Lemma: csm-revfill
[Gamma:j⊢]. ∀[A:{Gamma.𝕀 ⊢ _}]. ∀[cA:Gamma.𝕀 ⊢ Compositon(A)]. ∀[a1:{Gamma ⊢ _:(A)[1(𝕀)]}]. ∀[Delta:j⊢].
[s:Delta j⟶ Gamma].
  ((revfill(Gamma;cA;a1))s+ revfill(Delta;(cA)s+;(a1)s) ∈ {Delta.𝕀 ⊢ _:(A)s+})

Lemma: revfill-1
[Gamma:j⊢]. ∀[A:{Gamma.𝕀 ⊢ _}]. ∀[cA:Gamma.𝕀 ⊢ Compositon(A)]. ∀[a1:{Gamma ⊢ _:(A)[1(𝕀)]}].
  ((revfill(Gamma;cA;a1))[1(𝕀)] a1 ∈ {Gamma ⊢ _:(A)[1(𝕀)]})

Definition: filling_term
fill cA [phi ⊢→ u] a0 ==  fill cop-to-cfun(cA) [phi ⊢→ u] a0

Lemma: filling_term_wf
[Gamma:j⊢]. ∀[phi:{Gamma ⊢ _:𝔽}]. ∀[A:{Gamma.𝕀 ⊢ _}]. ∀[cA:Gamma.𝕀 ⊢ CompOp(A)]. ∀[u:{Gamma.𝕀(phi)p ⊢ _:A}].
[a0:{Gamma ⊢ _:(A)[0(𝕀)][phi |⟶ u[0]]}].
  (fill cA [phi ⊢→ u] a0 ∈ {Gamma.𝕀 ⊢ _:A[(phi)p |⟶ u]})

Lemma: csm-filling_term
[Gamma:j⊢]. ∀[phi:{Gamma ⊢ _:𝔽}]. ∀[A:{Gamma.𝕀 ⊢ _}]. ∀[cA:Gamma.𝕀 ⊢ CompOp(A)]. ∀[u:{Gamma.𝕀(phi)p ⊢ _:A}].
[a0:{Gamma ⊢ _:(A)[0(𝕀)][phi |⟶ u[0]]}]. ∀[Delta:j⊢]. ∀[s:Delta j⟶ Gamma].
  ((fill cA [phi ⊢→ u] a0)s+ fill (cA)s+ [(phi)s ⊢→ (u)s+] (a0)s ∈ {Delta.𝕀 ⊢ _:(A)s+[((phi)s)p |⟶ (u)s+]})

Lemma: filling_term_0
[H:j⊢]. ∀[phi:{H ⊢ _:𝔽}]. ∀[T:{H.𝕀 ⊢ _}]. ∀[u:{H.𝕀(phi)p ⊢ _:T}]. ∀[a0:{H ⊢ _:(T)[0(𝕀)][phi |⟶ u[0]]}].
[cT:H.𝕀 ⊢ CompOp(T)].
  ((fill cT [phi ⊢→ u] a0)[0(𝕀)] a0 ∈ {H ⊢ _:(T)[0(𝕀)]})

Lemma: filling_term_1
[H:j⊢]. ∀[phi:{H ⊢ _:𝔽}]. ∀[T:{H.𝕀 ⊢ _}]. ∀[u:{H.𝕀(phi)p ⊢ _:T}]. ∀[a0:{H ⊢ _:(T)[0(𝕀)][phi |⟶ u[0]]}].
[cT:H.𝕀 ⊢ CompOp(T)].
  ((fill cT [phi ⊢→ u] a0)[1(𝕀)] comp cT [phi ⊢→ u] a0 ∈ {H ⊢ _:(T)[1(𝕀)]})

Definition: fill-type-up
fill-type-up(Gamma;A;cA) ==  (fill (cA)p+ [0(𝔽) ⊢→ discr(⋅)] q)swap-interval(Gamma;(A)[0(𝕀)]))

Lemma: fill-type-up_wf
[Gamma:j⊢]. ∀[A:{Gamma.𝕀 ⊢ _}]. ∀[cA:Gamma.𝕀 ⊢ CompOp(A)].
  (fill-type-up(Gamma;A;cA) ∈ {Gamma.𝕀 ⊢ _:(((A)[0(𝕀)])p ⟶ A)})

Lemma: fill-type-up-0
[Gamma:j⊢]. ∀[A:{Gamma.𝕀 ⊢ _}]. ∀[cA:Gamma.𝕀 ⊢ CompOp(A)]. ∀[u:{Gamma ⊢ _:(A)[0(𝕀)]}].
  ((app(fill-type-up(Gamma;A;cA); (u)p))[0(𝕀)] u ∈ {Gamma ⊢ _:(A)[0(𝕀)]})

Lemma: fill-type-up-1
[Gamma:j⊢]. ∀[A:{Gamma.𝕀 ⊢ _}]. ∀[cA:Gamma.𝕀 ⊢ CompOp(A)]. ∀[u:{Gamma ⊢ _:(A)[0(𝕀)]}].
  ((app(fill-type-up(Gamma;A;cA); (u)p))[1(𝕀)] app(transport-fun(Gamma;A;cA); u) ∈ {Gamma ⊢ _:(A)[1(𝕀)]})

Definition: fill-type-down
fill-type-down(Gamma;A;cA) ==  (fill-type-up(Gamma;(A)-;(cA)(p;1-(q))))(p;1-(q))

Lemma: fill-type-down_wf
[Gamma:j⊢]. ∀[A:{Gamma.𝕀 ⊢ _}]. ∀[cA:Gamma.𝕀 ⊢ CompOp(A)].
  (fill-type-down(Gamma;A;cA) ∈ {Gamma.𝕀 ⊢ _:(((A)[1(𝕀)])p ⟶ A)})

Lemma: fill-type-down-1
[Gamma:j⊢]. ∀[A:{Gamma.𝕀 ⊢ _}]. ∀[cA:Gamma.𝕀 ⊢ CompOp(A)]. ∀[u:{Gamma ⊢ _:(A)[1(𝕀)]}].
  ((app(fill-type-down(Gamma;A;cA); (u)p))[1(𝕀)] u ∈ {Gamma ⊢ _:(A)[1(𝕀)]})

Lemma: fill-type-down-0
[Gamma:j⊢]. ∀[A:{Gamma.𝕀 ⊢ _}]. ∀[cA:Gamma.𝕀 ⊢ CompOp(A)]. ∀[u:{Gamma ⊢ _:(A)[1(𝕀)]}].
  ((app(fill-type-down(Gamma;A;cA); (u)p))[0(𝕀)] app(rev-transport-fun(Gamma;A;cA); u) ∈ {Gamma ⊢ _:(A)[0(𝕀)]})

Definition: fillpath
fillpath(Gamma;A;cA;x;y;z) ==
  comp (cA)p+ [((q=0) ∨ (q=1)) ⊢→ [(q)p=0 ⊢→ (app(fill-type-down(Gamma;A;cA); (x)p))p+;
                                   (q)p=1 ⊢→ (app(fill-type-up(Gamma;A;cA); (y)p))p+]] z

Lemma: fillpath_wf
[Gamma:j⊢]. ∀[A:{Gamma.𝕀 ⊢ _}]. ∀[cA:Gamma.𝕀 ⊢ CompOp(A)]. ∀[x:{Gamma ⊢ _:(A)[1(𝕀)]}]. ∀[y:{Gamma ⊢ _:(A)[0(𝕀)]}].
[z:{Gamma.𝕀 ⊢ _:((A)[0(𝕀)])p}].
  (fillpath(Gamma;A;cA;x;y;z) ∈ {t:{Gamma.𝕀 ⊢ _:((A)[1(𝕀)])p}| 
                                 ((t)[0(𝕀)] x ∈ {Gamma ⊢ _:(A)[1(𝕀)]})
                                 ∧ ((t)[1(𝕀)] app(transport-fun(Gamma;A;cA); y) ∈ {Gamma ⊢ _:(A)[1(𝕀)]})} supposing 
     (((z)[0(𝕀)] app(rev-transport-fun(Gamma;A;cA); x) ∈ {Gamma ⊢ _:(A)[0(𝕀)]}) and 
     ((z)[1(𝕀)] y ∈ {Gamma ⊢ _:(A)[0(𝕀)]}))

Definition: fill-path
fill-path(Gamma;A;cA;x;y;z) ==  <>(fillpath(Gamma;A;cA;x;y;z))

Lemma: fill-path_wf
[Gamma:j⊢]. ∀[A:{Gamma.𝕀 ⊢ _}]. ∀[cA:Gamma.𝕀 ⊢ CompOp(A)]. ∀[x:{Gamma ⊢ _:(A)[1(𝕀)]}]. ∀[y:{Gamma ⊢ _:(A)[0(𝕀)]}].
[z:{Gamma.𝕀 ⊢ _:((A)[0(𝕀)])p}].
  (fill-path(Gamma;A;cA;x;y;z) ∈ {Gamma ⊢ _:(Path_(A)[1(𝕀)] app(transport-fun(Gamma;A;cA); y))}) supposing 
     (((z)[0(𝕀)] app(rev-transport-fun(Gamma;A;cA); x) ∈ {Gamma ⊢ _:(A)[0(𝕀)]}) and 
     ((z)[1(𝕀)] y ∈ {Gamma ⊢ _:(A)[0(𝕀)]}))

Lemma: fill-path_wf_const
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[cA:Gamma ⊢ CompOp(A)]. ∀[x,y:{Gamma ⊢ _:A}]. ∀[z:{Gamma.𝕀 ⊢ _:(A)p}].
  (fill-path(Gamma;(A)p;(cA)p;x;y;z) ∈ {Gamma ⊢ _:(Path_A app(transport-fun(Gamma;(A)p;(cA)p); y))}) supposing 
     (((z)[0(𝕀)] app(rev-transport-fun(Gamma;(A)p;(cA)p); x) ∈ {Gamma ⊢ _:((A)p)[0(𝕀)]}) and 
     ((z)[1(𝕀)] y ∈ {Gamma ⊢ _:((A)p)[0(𝕀)]}))

Definition: trans-const-path
trans-const-path(G;cA;a) ==  <>(rev_fill_term(G;(cA)p;0(𝔽);discr(⋅);a))

Lemma: trans-const-path_wf
[G:j⊢]. ∀[A:{G ⊢ _}]. ∀[cA:G +⊢ Compositon(A)]. ∀[a:{G ⊢ _:A}].
  (trans-const-path(G;cA;a) ∈ {G ⊢ _:(Path_A transprt-const(G;cA;a) a)})

Definition: comp-path
pth_a_b pth_b_c ==
  <>(comp ((cA)p)p [((q=0) ∨ (q=1)) ⊢→ ((path-eta(refl(a)))p+ ∨ (path-eta(pth_b_c))p+)] path-eta(pth_a_b))

Lemma: comp-path_wf
[G:j⊢]. ∀[A:{G ⊢ _}]. ∀[cA:G ⊢ Compositon(A)]. ∀[a,b,c:{G ⊢ _:A}]. ∀[pth_a_b:{G ⊢ _:(Path_A b)}].
[pth_b_c:{G ⊢ _:(Path_A c)}].
  (pth_a_b pth_b_c ∈ {G ⊢ _:(Path_A c)})

Definition: comp_path
pth_a_b pth_b_c ==  pth_a_b pth_b_c

Lemma: comp_path_wf
[G:j⊢]. ∀[A:{G ⊢ _}]. ∀[cA:G ⊢ Compositon(A)]. ∀[a,b,c:{G ⊢ _:A}]. ∀[pth_a_b:{G ⊢ _:(Path_A b)}].
[pth_b_c:{G ⊢ _:(Path_A c)}].
  (pth_a_b pth_b_c ∈ {G ⊢ _:(Path_A c)})

Definition: contractible-to-prop
contractible-to-prop(X;A;cA;c) ==  rev-path(X.A.(A)p;contr-path(((c)p)p;(q)p)) contr-path(((c)p)p;q)))

Lemma: contractible-to-prop_wf
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[cA:X +⊢ Compositon(A)]. ∀[c:{X ⊢ _:Contractible(A)}].
  (contractible-to-prop(X;A;cA;c) ∈ {X ⊢ _:isProp(A)})

Lemma: contractible-iff-inhabited-prop
X:j⊢. ∀A:{X ⊢ _}. ∀cA:X +⊢ Compositon(A).  ({X ⊢ _:Contractible(A)} ⇐⇒ {X ⊢ _:isProp(A)} ∧ {X ⊢ _:A})

Definition: sigma_comp
sigma_comp(cA;cB) ==
  λH,sigma,phi,u,a0. let fill (cA)sigma [phi ⊢→ u.1] a0.1 in
                      let comp ((cB)sigma+)[a] [phi ⊢→ u.2] a0.2 in
                      cubical-pair((a)[1(𝕀)];b)

Lemma: sigma_comp_wf
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[B:{X.A ⊢ _}]. ∀[cA:X ⊢ Compositon(A)]. ∀[cB:X.A +⊢ Compositon(B)].
  (sigma_comp(cA;cB) ∈ composition-function{j:l,i:l}(X;Σ B))

Lemma: sigma_comp-sq
[X,A,cA,cB:Top].
  (sigma_comp(cA;cB) ~ λH,sigma,phi,u,a0. let cA H.𝕀 x,x@0. (sigma (m x@0))) 
                                                  I,rho. phi (fst(rho)) ∨ dM-to-FL(I;¬(snd(rho)))) 
                                                  I,rho. if (phi (fst((m rho)))==1)
                                                          then fst((u (m rho)))
                                                          else fst((a0 (fst((m rho)))))
                                                          fi 
                                                  I,a. (fst((a0 (fst(a)))))) in
                                           let cB x,x@0. <sigma x@0, x@0>phi I,a. (snd((u a)))) 
                                                   I,a. (snd((a0 a)))) in
                                           λI,a@0. <I <a@0, 1>a@0>)

Lemma: sigma_comp_wf2
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[B:{X.A ⊢ _}]. ∀[cA:X ⊢ Compositon(A)]. ∀[cB:X.A +⊢ Compositon(B)].
  (sigma_comp(cA;cB) ∈ X ⊢ Compositon(Σ B))

Lemma: csm-sigma_comp
[X,A,cA,cB,H,tau:Top].  ((sigma_comp(cA;cB))tau sigma_comp((cA)tau;(cB)tau+))

Lemma: csm-sigma_comp2
[X,H,cA,cB,A,tau:Top].  ((sigma_comp(cA;cB))tau sigma_comp((cA)tau;(cB)tau+))

Lemma: csm-sigma_comp3
[H,A',X,cA,cB,A,tau:Top].  ((sigma_comp(cA;cB))tau sigma_comp((cA)tau;(cB)tau+))

Lemma: fst-transprt-sigma
[X:j⊢]. ∀[A:{X.𝕀 ⊢ _}]. ∀[B:{X.𝕀.A ⊢ _}]. ∀[cA:X.𝕀 +⊢ Compositon(A)]. ∀[cB:X.𝕀.A +⊢ Compositon(B)].
[pr:{X ⊢ _:(Σ B)[0(𝕀)]}].
  (transprt(X;sigma_comp(cA;cB);pr).1 transprt(X;cA;pr.1) ∈ {X ⊢ _:(A)[1(𝕀)]})

Lemma: fst-transprt-const-sigma
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[B:{X.A ⊢ _}]. ∀[cA:X +⊢ Compositon(A)]. ∀[cB:X.A +⊢ Compositon(B)]. ∀[pr:{X ⊢ _:Σ B}].
  (transprt-const(X;sigma_comp(cA;cB);pr).1 transprt-const(X;cA;pr.1) ∈ {X ⊢ _:A})

Definition: pi_comp
pi_comp(X;A;cA;cB) ==
  λH,sigma,phi,u,a0. let revfill(H.((A)sigma)[1(𝕀)];((cA)sigma)p+;q) in
                         comp (cB)(sigma p+;v) [(phi)p ⊢→ app((u)p+; v)] app((a0)p; (v)[0(𝕀)]))

Lemma: pi_comp_wf
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[B:{X.A ⊢ _}]. ∀[cA:X +⊢ Compositon(A)]. ∀[cB:X.A +⊢ Compositon(B)].
  (pi_comp(X;A;cA;cB) ∈ composition-function{j:l,i:l}(X;ΠB))

Lemma: pi_comp-sq
[X,A,cA,cB:Top].
  (pi_comp(X;A;cA;cB) ~ λH,sigma,phi,u,a0. let = λI,a. (cA H.((A)sigma)[1(𝕀)].𝕀 
                                                          x,x@0. (sigma x <fst(fst((m x@0))), ¬(snd((m x@0)))>)) 
                                                          I,rho. 0 ∨ dM-to-FL(I;¬(snd(rho)))) 
                                                          I,rho. (snd(fst((m rho))))) 
                                                          I,a. (snd(fst(a)))) 
                                                          
                                                          <fst(a), ¬(snd(a))>in
                                               cB H.((A)sigma)[1(𝕀)] x,x@0. <sigma x <fst(fst(x@0)), snd(x@0)>\000Cx@0>
                                                 I,a. (phi (fst(a)))) 
                                                 I,a. (u I <fst(fst(a)), snd(a)> (v a))) 
                                                 I,a. (a0 (fst(a)) (v I <a, 0>)))))

Lemma: csm-pi_comp
[X,Y,tau,A,cA,cB:Top].  ((pi_comp(X;A;cA;cB))tau pi_comp(Y;(A)tau;(cA)tau;(cB)tau+))

Lemma: pi_comp_wf2
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[B:{X.A ⊢ _}]. ∀[cA:X +⊢ Compositon(A)]. ∀[cB:X.A +⊢ Compositon(B)].
  (pi_comp(X;A;cA;cB) ∈ X ⊢ Compositon(ΠB))

Lemma: pi_comp_wf_fun
[X:j⊢]. ∀[A,B:{X ⊢ _}]. ∀[cA:X +⊢ Compositon(A)]. ∀[cB:X +⊢ Compositon(B)].
  (pi_comp(X;A;cA;(cB)p) ∈ X ⊢ Compositon((A ⟶ B)))

Definition: pathtype-comp
pathtype-comp(G;A;cA) ==  λH,sigma,phi,u,a0. <>comp ((cA)sigma)p+ [(phi)p ⊢→ (u)p+ (q)p] (a0)p q

Lemma: pathtype-comp_wf
[G:j⊢]. ∀[A:{G ⊢ _}]. ∀[cA:G ⊢ Compositon(A)].  (pathtype-comp(G;A;cA) ∈ G ⊢ Compositon(Path(A)))

Lemma: csm-pathtype-comp
[G,A,cA,H,tau:Top].  ((pathtype-comp(G;A;cA))tau pathtype-comp(H;(A)tau;(cA)tau))

Definition: path-term
path-term(phi;w;a;b;r) ==  (w r ∨ (a ∨ b))

Lemma: path-term_wf
X:j⊢. ∀psi:{X ⊢ _:𝔽}.
  ∀[r:{X ⊢ _:𝕀}]
    ∀T:{X ⊢ _}. ∀a,b:{X ⊢ _:T}. ∀w:{X, psi ⊢ _:(Path_T b)}.
      (path-term(psi;w;a;b;r) ∈ {X, (psi ∨ ((r=0) ∨ (r=1))) ⊢ _:T})

Lemma: path-term-case1
X:j⊢. ∀psi:{X ⊢ _:𝔽}.
  ∀[r:{X ⊢ _:𝕀}]
    ∀T:{X ⊢ _}. ∀a,b:{X ⊢ _:T}. ∀w:{X, psi ⊢ _:(Path_T b)}.  (path-term(psi;w;a;b;r) r ∈ {X, psi ⊢ _:T})

Lemma: path-term-equal
X:j⊢. ∀psi:{X ⊢ _:𝔽}.
  ∀[r:{X ⊢ _:𝕀}]
    ∀T:{X ⊢ _}. ∀a,b:{X ⊢ _:T}. ∀w:{X, psi ⊢ _:(Path_T b)}.
      ∀[z:{X ⊢ _:T}]
        path-term(psi;w;a;b;r) z ∈ {X, (psi ∨ ((r=0) ∨ (r=1))) ⊢ _:T} 
        supposing (w z ∈ {X, psi ⊢ _:T}) ∧ (a z ∈ {X, (r=0) ⊢ _:T}) ∧ (b z ∈ {X, (r=1) ⊢ _:T})

Lemma: path-term-0
X:j⊢. ∀psi:{X ⊢ _:𝔽}. ∀T:{X ⊢ _}. ∀a,b:{X ⊢ _:T}. ∀w:{X, psi ⊢ _:(Path_T b)}.
  (path-term(psi;w;a;b;0(𝕀)) a ∈ {X ⊢ _:T})

Lemma: path-term-1
X:j⊢. ∀psi:{X ⊢ _:𝔽}. ∀T:{X ⊢ _}. ∀a,b:{X ⊢ _:T}. ∀w:{X, psi ⊢ _:(Path_T b)}.
  (path-term(psi;w;a;b;1(𝕀)) b ∈ {X ⊢ _:T})

Lemma: csm-path-term
[phi,r,s,a,b,w:Top].  ((path-term(phi;w;a;b;r))s path-term((phi)s;(w)s;(a)s;(b)s;(r)s))

Definition: path_term
path_term(phi; w; a; b; r) ==  path-term((phi)p;w;a;b;(r)p)

Lemma: path_term_wf
X:j⊢. ∀psi:{X ⊢ _:𝔽}.
  ∀[r:{X ⊢ _:𝕀}]
    ∀T:{X.𝕀 ⊢ _}. ∀a,b:{X.𝕀 ⊢ _:T}. ∀w:{X, psi.𝕀 ⊢ _:(Path_T b)}.
      (path_term(psi; w; a; b; r) ∈ {X.𝕀((psi)p ∨ (((r=0))p ∨ ((r=1))p)) ⊢ _:T})

Lemma: csm-path_term
[psi,r,s,a,b,w:Top].  ((path_term(psi; w; a; b; r))s path-term(((psi)p)s;(w)s;(a)s;(b)s;((r)p)s))

Lemma: path_term-0
[H,A,psi,r,a,b,w:Top].
  ((path_term(psi; w; a; b; r))[0(𝕀)] path-term((psi)1(H.A);(w)[0(𝕀)];(a)[0(𝕀)];(b)[0(𝕀)];(r)1(H.A)))

Lemma: path_term-1
[H,A,psi,r,a,b,w:Top].
  ((path_term(psi; w; a; b; r))[1(𝕀)] path-term((psi)1(H.A);(w)[1(𝕀)];(a)[1(𝕀)];(b)[1(𝕀)];(r)1(H.A)))

Definition: path_comp
path_comp(G;A;a;b;
          cA) ==
  λH,sigma,phi,u,a0. <>(comp ((cA)sigma)p+ [((phi)p ∨ ((q=0) ∨ (q=1))) ⊢→ path_term((phi)p;
                                                                                    (u)p+;
                                                                                    ((a)sigma)p+;
                                                                                    ((b)sigma)p+;
                                                                                    q)] (a0)p q)

Lemma: path_comp_wf
[G:j⊢]. ∀[A:{G ⊢ _}]. ∀[a,b:{G ⊢ _:A}]. ∀[cA:G ⊢ Compositon(A)].
  (path_comp(G;A;a;b;
             cA) ∈ G ⊢ Compositon((Path_A b)))

Lemma: path-comp-exists
G:j⊢. ∀A:{G ⊢ _}. ∀a,b:{G ⊢ _:A}.  (G ⊢ CompOp(A)  G ⊢ CompOp((Path_A b)))

Lemma: csm-path_comp
[G,H,A,a,b,cA,tau:Top].  ((path_comp(G;A;a;b;cA))tau path_comp(H;(A)tau;(a)tau;(b)tau;(cA)tau))

Lemma: cubical-pi-comp-structure
X:j⊢. ∀A:{X ⊢ _}. ∀B:{X.A ⊢ _}.  (X ⊢ Compositon(A)  X.A +⊢ Compositon(B)  X ⊢ Compositon(ΠB))

Lemma: transform-comp-structure
Gamma,Delta:j⊢. ∀tau:Delta j⟶ Gamma. ∀A:{Gamma ⊢ _}.  (Gamma ⊢ Compositon(A)  Delta ⊢ Compositon((A)tau))

Definition: fiber-comp
fiber-comp(X;T;A;w;a;cT;cA) ==  sigma_comp(cT;path_comp(X.T;(A)p;(a)p;app((w)p; q);(cA)p))

Lemma: fiber-comp_wf
[X:j⊢]. ∀[T,A:{X ⊢ _}]. ∀[w:{X ⊢ _:(T ⟶ A)}]. ∀[a:{X ⊢ _:A}]. ∀[cT:X ⊢ Compositon(T)]. ∀[cA:X +⊢ Compositon(A)].
  (fiber-comp(X;T;A;w;a;cT;cA) ∈ X ⊢ Compositon(Fiber(w;a)))

Lemma: fiber-comp-exists
X:j⊢. ∀T,A:{X ⊢ _}. ∀w:{X ⊢ _:(T ⟶ A)}. ∀a:{X ⊢ _:A}.  (X ⊢ CompOp(T)  X ⊢ CompOp(A)  X ⊢ CompOp(Fiber(w;a)))

Lemma: csm-fiber-comp-sq
[G,A,T,a,cA,cT,H,s,f:Top].  ((fiber-comp(G;T;A;f;a;cT;cA))s fiber-comp(H;(T)s;(A)s;(f)s;(a)s;(cT)s;(cA)s))

Lemma: fiber-comp-subset
[X,T,A,w,a,cT,cA,phi:Top].  (fiber-comp(X, phi;T;A;w;a;cT;cA) fiber-comp(X;T;A;w;a;cT;cA))

Lemma: csm-fiber-comp
[G:j⊢]. ∀[A,T:{G ⊢ _}]. ∀[a:{G ⊢ _:A}]. ∀[cA:G +⊢ Compositon(A)]. ∀[cT:G ⊢ Compositon(T)]. ∀[H:j⊢]. ∀[s:H j⟶ G].
[f:{G ⊢ _:(T ⟶ A)}].
  ((fiber-comp(G;T;A;f;a;cT;cA))s fiber-comp(H;(T)s;(A)s;(f)s;(a)s;(cT)s;(cA)s) ∈ H ⊢ Compositon(Fiber((f)s;(a)s)))

Lemma: fiber-member-transprt-const-fiber-comp
[X:j⊢]. ∀[T,A:{X ⊢ _}]. ∀[w:{X ⊢ _:(T ⟶ A)}]. ∀[a:{X ⊢ _:A}]. ∀[pr:{X ⊢ _:Fiber(w;a)}]. ∀[cT:X +⊢ Compositon(T)].
[cA:X +⊢ Compositon(A)].
  (fiber-member(transprt-const(X;fiber-comp(X;T;A;w;a;cT;cA);pr)) transprt-const(X;cT;pr.1) ∈ {X ⊢ _:T})

Definition: contractible_comp
contractible_comp(X;A;cA) ==  sigma_comp(cA;pi_comp(X.A;(A)p;(cA)p;path_comp(X.A.(A)p;((A)p)p;(q)p;q;((cA)p)p)))

Lemma: contractible_comp_wf
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[cA:X +⊢ Compositon(A)].  (contractible_comp(X;A;cA) ∈ X ⊢ Compositon(Contractible(A)))

Lemma: contractible_comp-sq
[X,A,cA:Top].
  (contractible_comp(X;A;cA) 
  ~ λH,sigma,phi,u,a0.
                      let cA H.𝕀 x,x@0. (sigma (m x@0))) I,rho. phi (fst(rho)) ∨ dM-to-FL(I;¬(snd(rho)))) 
                              I,rho. if (phi (fst((m rho)))==1)
                                      then fst((u (m rho)))
                                      else fst((a0 (fst((m rho)))))
                                      fi 
                              I,a. (fst((a0 (fst(a)))))) in
                       let let = λI,a@0. (cA H.((A)sigma)[1(𝕀)].𝕀 
                                                x,x@0. (sigma x <fst(fst((m x@0))), ¬(snd((m x@0)))>)) 
                                                I,rho. 0 ∨ dM-to-FL(I;¬(snd(rho)))) 
                                                I,rho. if (0==1) then ⋅ else snd(fst((m rho))) fi 
                                                I,a. (snd(fst(a)))) 
                                                
                                                <fst(a@0), ¬(snd(a@0))>in
                                   λI,a@0,J,f,u@0,J@0,f@0,u@1.
                                                              (cA H.((A)sigma)[1(𝕀)].𝕀 
                                                               x,x@0. (sigma x <fst(fst(fst(x@0))), snd(x@0)>)) 
                                                               I,rho.
                                                                       phi (fst(fst(rho))) ∨ dM-to-FL(I;¬(...)) ∨ ...)\000C 
                                                               I,rho. if (phi (fst(fst(fst(rho))))==1)
                                                                          then (snd((u I <fst(fst(fst(rho))), snd(rho)>)\000C)) 
                                                                               
                                                                               
                                                                               (v I <fst(fst(rho)), snd(rho)>
                                                                               
                                                                               
                                                                               (snd(fst(rho)))
                                                                       if (dM-to-FL(I;¬(snd(fst(rho))))==1)
                                                                         then I <fst(fst(fst(rho))), snd(rho)>
                                                                       else I <fst(fst(rho)), snd(rho)>
                                                                       fi 
                                                               I,a. ((snd((a0 (fst(fst(a)))))) (v I <fst(a), 0>)\000C 
                                                                       
                                                                       (snd(a)))) 
                                                               J@0 
                                                               <f@0((f(a@0);u@0)), u@1>in
                       cubical-pair(λI,a@0. (a I <a@0, 1(𝕀a@0>);b))

Lemma: contractible_comp-apply-sq
[X,A,cA,B,a,b,c,d,J,x:Top].
  (contractible_comp(X;A;cA) x,y. a[x;y]) x,y. b[x;y]) x,y. c[x;y]) x,y. d[x;y]) 
  ~ <cA <λI.(alpha:B(I) × 𝕀(alpha)), λI,J,f,p. <f(fst(p)), (snd(p) fst(p) f)>> x,x@0. a[x;m x@0]) 
     I,rho. b[I;fst(rho)] ∨ dM-to-FL(I;¬(snd(rho)))) 
     I,rho. if (b[I;fst((m rho))]==1) then fst(c[I;m rho]) else fst(d[I;fst((m rho))]) fi 
     I,a. (fst(d[I;fst(a)]))) 
     
     <x, 1>
    , λJ@0,f,u@0,J@1,f@0,u@1. (cA 
                               <λI.(alpha:alpha:B(I) × ((A)λx,y. a[x;y])[λI,rho. 1](alpha) × 𝕀(alpha))
                               , λI,J,f,p. <<f(fst(fst(p))), (snd(fst(p)) fst(fst(p)) f)>(snd(p) fst(p) f)>
                               > 
                               x,x@0. a[x;<fst(fst(fst(x@0))), snd(x@0)>]) 
                               I,rho. b[I;fst(fst(rho))] ∨ dM-to-FL(I;¬(snd(rho))) ∨ dM-to-FL(I;snd(rho))) 
                               I,rho. if (b[I;fst(fst(fst(rho)))]==1)
                                          then (snd(c[I;<fst(fst(fst(rho))), snd(rho)>])) 
                                               (cA 
                                                <λI.(alpha:alpha:B(I) × ((A)λx,y. a[x;y])[λI,rho. 1](alpha) × 𝕀(alpha))
                                                , λI,J,f,p. <<f(fst(fst(p))), (snd(fst(p)) fst(fst(p)) f)>
                                                            (snd(p) fst(p) f)
                                                            >
                                                > 
                                                x,x@0. a[x;<fst(fst((m x@0))), ¬(snd((m x@0)))>]) 
                                                I,rho. 0 ∨ dM-to-FL(I;¬(snd(rho)))) 
                                                I,rho. (snd(fst((m rho))))) 
                                                I,a. (snd(fst(a)))) 
                                                
                                                <fst(fst(rho)), ¬(snd(rho))>
                                               
                                               
                                               (snd(fst(rho)))
                                       if (dM-to-FL(I;¬(snd(fst(rho))))==1)
                                         then cA <λI.(alpha:B(I) × 𝕀(alpha)), λI,J,f,p. <f(fst(p)), (snd(p) fst(p) f)>> 
                                              x,x@0. a[x;m x@0]) 
                                              I,rho. b[I;fst(rho)] ∨ dM-to-FL(I;¬(snd(rho)))) 
                                              I,rho. if (b[I;fst((m rho))]==1)
                                                      then fst(c[I;m rho])
                                                      else fst(d[I;fst((m rho))])
                                                      fi 
                                              I,a. (fst(d[I;fst(a)]))) 
                                              
                                              <fst(fst(fst(rho))), snd(rho)>
                                       else cA 
                                            <λI.(alpha:alpha:B(I) × ((A)λx,y. a[x;y])[λI,rho. 1](alpha) × 𝕀(alpha))
                                            , λI,J,f,p. <<f(fst(fst(p))), (snd(fst(p)) fst(fst(p)) f)>
                                                        (snd(p) fst(p) f)
                                                        >
                                            > 
                                            x,x@0. a[x;<fst(fst((m x@0))), ¬(snd((m x@0)))>]) 
                                            I,rho. 0 ∨ dM-to-FL(I;¬(snd(rho)))) 
                                            I,rho. (snd(fst((m rho))))) 
                                            I,a. (snd(fst(a)))) 
                                            
                                            <fst(fst(rho)), ¬(snd(rho))>
                                       fi 
                               I,a@0. ((snd(d[I;fst(fst(a@0))])) 
                                         (cA 
                                          <λI.(alpha:alpha:B(I) × ((A)λx,y. a[x;y])[λI,rho. 1](alpha) × 𝕀(alpha))
                                          , λI,J,f,p. <<f(fst(fst(p))), (snd(fst(p)) fst(fst(p)) f)>(snd(p) fst(p) f)>
                                          > 
                                          x,x@0. a[x;<fst(fst((m x@0))), ¬(snd((m x@0)))>]) 
                                          I,rho. 0 ∨ dM-to-FL(I;¬(snd(rho)))) 
                                          I,rho. (snd(fst((m rho))))) 
                                          I,a. (snd(fst(a)))) 
                                          
                                          <fst(a@0), ¬(0)>
                                         
                                         
                                         (snd(a@0)))) 
                               J@1 
                               <<f@0(f(x)), (u@0 f(x) f@0)>u@1>)
    >)

Lemma: csm-contractible_comp
[X,H,cA,A,tau:Top].  ((contractible_comp(X;A;cA))tau contractible_comp(H;(A)tau;(cA)tau))

Definition: contractible-comp
contractible-comp(X;A;cA) ==  cfun-to-cop(X;Contractible(A);contractible_comp(X;A;cop-to-cfun(cA)))

Lemma: contractible-comp_wf
[X:j⊢]. ∀[A:{X ⊢ _}]. ∀[cA:X ⊢ CompOp(A)].  (contractible-comp(X;A;cA) ∈ X ⊢ CompOp(Contractible(A)))

Lemma: contractible-comp-exists
X:j⊢. ∀A:{X ⊢ _}.  (X ⊢ CompOp(A)  X ⊢ CompOp(Contractible(A)))

Definition: compatible-composition
compatible-composition{j:l, i:l}(Gamma; phi; psi; A; B; cA; cB) ==
  ∀H:j⊢. ∀sigma:H.𝕀 j⟶ Gamma, (phi ∧ psi). ∀chi:{H ⊢ _:𝔽}. ∀u:{H, chi.𝕀 ⊢ _:(A)sigma}.
  ∀a0:{H ⊢ _:((A)sigma)[0(𝕀)][chi |⟶ (u)[0(𝕀)]]}.
    ((cB sigma chi a0) (cA sigma chi a0) ∈ {H ⊢ _:((A)sigma)[1(𝕀)]})

Lemma: compatible-composition_wf
[Gamma:j⊢]. ∀[phi,psi:{Gamma ⊢ _:𝔽}]. ∀[A:{Gamma, phi ⊢ _}]. ∀[B:{Gamma, psi ⊢ _}]. ∀[cA:Gamma, phi ⊢ Compositon(A)].
[cB:Gamma, psi ⊢ Compositon(B)].
  compatible-composition{j:l, i:l}(Gamma; phi; psi; A; B; cA; cB) ∈ ℙ{[i j'']} supposing Gamma, (phi ∧ psi) ⊢ B

Lemma: compatible-composition-disjoint
[Gamma:j⊢]. ∀[phi,psi:{Gamma ⊢ _:𝔽}]. ∀[A:{Gamma, phi ⊢ _}]. ∀[B:{Gamma, psi ⊢ _}]. ∀[cA:Gamma, phi ⊢ Compositon(A)].
[cB:Gamma, psi ⊢ Compositon(B)].
  compatible-composition{j:l, i:l}(Gamma; phi; psi; A; B; cA; cB) supposing Gamma ⊢ ((phi ∧ psi)  0(𝔽))

Definition: case-type-comp
case-type-comp(G; phi; psi; A; B; cA; cB) ==
  λH,sigma,chi,u,a0. (cA H, (∀ (phi)sigma) sigma chi a0 ∨ cB H, (∀ (psi)sigma) sigma chi a0)

Lemma: case-type-comp_wf1
[Gamma:j⊢]. ∀[phi,psi:{Gamma ⊢ _:𝔽}]. ∀[A:{Gamma, phi ⊢ _}]. ∀[B:{Gamma, psi ⊢ _}]. ∀[cA:Gamma, phi ⊢ Compositon(A)].
[cB:Gamma, psi ⊢ Compositon(B)].
  (case-type-comp(Gamma; phi; psi; A; B; cA; cB)
   ∈ composition-function{j:l,i:l}(Gamma, (phi ∨ psi);(if phi then else B))) supposing 
     (compatible-composition{j:l, i:l}(Gamma; phi; psi; A; B; cA; cB) and 
     Gamma, (phi ∧ psi) ⊢ B)

Lemma: case-type-comp_wf
[Gamma:j⊢]. ∀[phi,psi:{Gamma ⊢ _:𝔽}]. ∀[A:{Gamma, phi ⊢ _}]. ∀[B:{Gamma, psi ⊢ _}]. ∀[cA:Gamma, phi ⊢ Compositon(A)].
[cB:Gamma, psi ⊢ Compositon(B)].
  (case-type-comp(Gamma; phi; psi; A; B; cA; cB) ∈ Gamma, (phi ∨ psi) ⊢ Compositon((if phi then else B))) supposing 
     (compatible-composition{j:l, i:l}(Gamma; phi; psi; A; B; cA; cB) and 
     Gamma, (phi ∧ psi) ⊢ B)

Lemma: case-type-comp-true-false
[Gamma:j⊢]. ∀[phi,psi:{Gamma ⊢ _:𝔽}].
  (∀[A:{Gamma ⊢ _}]. ∀[cA:Gamma ⊢ Compositon(A)]. ∀[B:{Gamma, psi ⊢ _}]. ∀[cB:Gamma, psi ⊢ Compositon(B)].
     (case-type-comp(Gamma; phi; psi; A; B; cA; cB) cA ∈ Gamma ⊢ Compositon(A))) supposing 
     (Gamma ⊢ (1(𝔽 phi) and 
     Gamma ⊢ (psi  0(𝔽)))

Lemma: case-type-comp-false-true
[Gamma:j⊢]. ∀[phi,psi:{Gamma ⊢ _:𝔽}].
  (∀[B:{Gamma ⊢ _}]. ∀[cB:Gamma ⊢ Compositon(B)]. ∀[A:{Gamma, phi ⊢ _}]. ∀[cA:Gamma, phi ⊢ Compositon(A)].
     (case-type-comp(Gamma; phi; psi; A; B; cA; cB) cB ∈ Gamma ⊢ Compositon(B))) supposing 
     (Gamma ⊢ (1(𝔽 psi) and 
     (phi 0(𝔽) ∈ {Gamma ⊢ _:𝔽}))

Lemma: csm-case-type-comp
[H,Gamma,tau,phi,psi,A,B,cA,cB:Top].
  ((case-type-comp(Gamma; phi; psi; A; B; cA; cB))tau case-type-comp(H;
                                                                       (phi)tau;
                                                                       (psi)tau;
                                                                       (A)tau;
                                                                       (B)tau;
                                                                       (cA)tau;
                                                                       (cB)tau))

Lemma: case-type-comp-disjoint
[Gamma:j⊢]. ∀[phi,psi:{Gamma ⊢ _:𝔽}]. ∀[A:{Gamma, phi ⊢ _}]. ∀[B:{Gamma, psi ⊢ _}]. ∀[cA:Gamma, phi ⊢ Compositon(A)].
[cB:Gamma, psi ⊢ Compositon(B)].
  case-type-comp(Gamma; phi; psi; A; B; cA; cB) ∈ Gamma, (phi ∨ psi) ⊢ Compositon((if phi then else B)) 
  supposing Gamma ⊢ ((phi ∧ psi)  0(𝔽))

Lemma: case-type-comp-partition
[Gamma:j⊢]. ∀[phi,psi:{Gamma ⊢ _:𝔽}]. ∀[A:{Gamma, phi ⊢ _}]. ∀[B:{Gamma, psi ⊢ _}]. ∀[cA:Gamma, phi ⊢ Compositon(A)].
[cB:Gamma, psi ⊢ Compositon(B)].
  case-type-comp(Gamma; phi; psi; A; B; cA; cB) ∈ Gamma ⊢ Compositon((if phi then else B)) 
  supposing Gamma ⊢ ((phi ∧ psi)  0(𝔽)) ∧ Gamma ⊢ (1(𝔽 (phi ∨ psi))

Definition: fibrant-type
FibrantType(X) ==  A:{X ⊢ _} × X ⊢ CompOp(A)

Lemma: fibrant-type_wf
[X:j⊢]. (FibrantType(X) ∈ 𝕌{[i' j']})

Lemma: fibrant-type_wf_formal-cube
[I:fset(ℕ)]. (FibrantType(formal-cube(I)) ∈ 𝕌')

Lemma: fibrant-type-cumulativity
[X:j⊢]. (FibrantType(X) ⊆fibrant-type{i':l}(X))

Definition: csm-fibrant-type
csm-fibrant-type(G;H;s;FT) ==  let A,cA FT in <(A)s, (cA)s>

Lemma: csm-fibrant-type_wf
[G,H:j⊢]. ∀[s:H j⟶ G]. ∀[FT:FibrantType(G)].  (csm-fibrant-type(G;H;s;FT) ∈ FibrantType(H))

Lemma: csm-fibrant-type-id
[G:j⊢]. ∀[FT:FibrantType(G)]. ∀[tau:G j⟶ G].
  csm-fibrant-type(G;G;tau;FT) FT ∈ FibrantType(G) supposing tau 1(G) ∈ j⟶ G

Lemma: csm-fibrant-comp
[X,Y,Z:j⊢]. ∀[g:Z j⟶ Y]. ∀[f:Y j⟶ X]. ∀[FT:FibrantType(X)].
  (csm-fibrant-type(X;Z;f g;FT) csm-fibrant-type(Y;Z;g;csm-fibrant-type(X;Y;f;FT)) ∈ FibrantType(Z))

Definition: discrete-comp
discrete-comp(G;T) ==  λI,i,rho,phi,u,a0. a0

Lemma: discrete-comp_wf
[G:j⊢]. ∀[T:Type].  (discrete-comp(G;T) ∈ G ⊢ CompOp(discr(T)))

Definition: discrete_comp
discrete_comp(G;T) ==  λH,sigma,phi,u,a0. a0

Lemma: discrete_comp_wf
[G:j⊢]. ∀[T:Type].  (discrete_comp(G;T) ∈ G ⊢ Compositon(discr(T)))

Definition: extension-fun
extension-fun{i:l}(Gamma;A) ==
  H:CubicalSet ⟶ sigma:H ⟶ Gamma ⟶ phi:{H ⊢ _:𝔽} ⟶ u:{H, phi ⊢ _:(A)sigma} ⟶ {H ⊢ _:(A)sigma[phi |⟶ u]}

Lemma: extension-fun_wf
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}].  (extension-fun{j:l}(Gamma; A) ∈ 𝕌{[i j'']})

Definition: uniform-extension-fun
uniform-extension-fun{i:l}(Gamma;A;ext) ==
  ∀H,K:⊢. ∀tau:K ⟶ H. ∀sigma:H ⟶ Gamma. ∀phi:{H ⊢ _:𝔽}. ∀u:{H, phi ⊢ _:(A)sigma}.
    ((ext sigma phi u)tau (ext sigma tau (phi)tau (u)tau) ∈ {K ⊢ _:((A)sigma)tau[(phi)tau |⟶ (u)tau]})

Lemma: uniform-extension-fun_wf
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[ext:extension-fun{j:l}(Gamma; A)].
  (uniform-extension-fun{j:l}(Gamma; A; ext) ∈ ℙ{[i j'']})

Definition: uniform-extend
uniform-extend{i:l}(Gamma; A) ==  {ext:extension-fun{i:l}(Gamma;A)| uniform-extension-fun{i:l}(Gamma;A;ext)} 

Lemma: uniform-extend_wf
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}].  (Gamma ⊢ Extension(A) ∈ 𝕌{[i j'']})

Definition: extend-to-contraction
extend-to-contraction(Gamma;A;ext) ==
  contr-witness(Gamma;ext Gamma 1(Gamma) 0(𝔽discr(⋅);<>(ext Gamma.A.𝕀 (q=1) (q)p))

Lemma: extend-to-contraction_wf
Gamma:j⊢. ∀A:{Gamma ⊢ _}. ∀ext:Gamma +⊢ Extension(A).
  (extend-to-contraction(Gamma;A;ext) ∈ {Gamma ⊢ _:Contractible(A)})

Definition: contraction-to-extend
contraction-to-extend(Gamma;A;cA;ctr) ==
  λH,sigma,phi,u. comp ((cA)sigma)p [phi ⊢→ path-eta(contr-path((ctr)sigma;u))] contr-center((ctr)sigma)

Lemma: contraction-to-extend_wf
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[cA:Gamma ⊢ Compositon(A)]. ∀[ctr:{Gamma ⊢ _:Contractible(A)}].
  (contraction-to-extend(Gamma;A;cA;ctr) ∈ Gamma ⊢ Extension(A))

Lemma: uniform-extend-contractible
Gamma:j⊢. ∀A:{Gamma ⊢ _}. ∀ext:Gamma +⊢ Extension(A).  {Gamma ⊢ _:Contractible(A)}

Lemma: contractible-uniform-extend
Gamma:j⊢. ∀A:{Gamma ⊢ _}. ∀cA:Gamma ⊢ Compositon(A). ∀ctr:{Gamma ⊢ _:Contractible(A)}.  Gamma ⊢ Extension(A)

Definition: extend-face-term
This construction extends partial term u ∈ {I,phi ⊢ _:𝔽to a
total formula in 𝔽(I). 

For each face, fc (irr_face(I;as;bs)),
in the components of phi, there is "free-est" morphism, 
 g=(irr-face-morph(I;as;bs)), that makes (fc)g 1.
This makes (fc ∧ u(g)) well-defined formula in 𝔽(I).
The join of all of these formulas is the extension.

It satisfies these additional properties:
extend-face-term-property
extend-face-term-le
extend-face-term-unique
extend-face-term-morph

We use it to define the composition operation for ⌜𝔽⌝
face_comp.
face_comp_wf

extend-face-term(I;phi;u) ==
  \/(λpr.let as,bs pr 
         in irr_face(I;as;bs) ∧ u(irr-face-morph(I;as;bs))"(face_lattice_components(I;phi)))

Lemma: extend-face-term_wf
[I:fset(ℕ)]. ∀[phi:𝔽(I)]. ∀[u:{I,phi ⊢ _:𝔽}].  (extend-face-term(I;phi;u) ∈ 𝔽(I))

Lemma: extend-face-term-property
[I:fset(ℕ)]. ∀[phi:𝔽(I)]. ∀[u:{I,phi ⊢ _:𝔽}]. ∀[J:fset(ℕ)]. ∀[g:I,phi(J)].
  ((extend-face-term(I;phi;u))<g> u(g) ∈ Point(face_lattice(J)))

Lemma: extend-face-term-le
[I:fset(ℕ)]. ∀[phi:𝔽(I)]. ∀[u:{I,phi ⊢ _:𝔽}].  extend-face-term(I;phi;u) ≤ phi

Lemma: extend-face-term-uniqueness
[I:fset(ℕ)]. ∀[phi:Point(face_lattice(I))]. ∀[u:{I,phi ⊢ _:𝔽}]. ∀[a,b:Point(face_lattice(I))].
  b ∈ Point(face_lattice(I)) 
  supposing a ≤ phi
  ∧ b ≤ phi
  ∧ (∀[g:{f:I ⟶ I| (phi f) 1} ]. ((a)<g> u(g) ∈ Point(face_lattice(I))))
  ∧ (∀[g:{f:I ⟶ I| (phi f) 1} ]. ((b)<g> u(g) ∈ Point(face_lattice(I))))

Lemma: extend-face-term-unique
[I:fset(ℕ)]. ∀[phi:Point(face_lattice(I))]. ∀[u:{I,phi ⊢ _:𝔽}]. ∀[a:Point(face_lattice(I))].
  extend-face-term(I;phi;u) ∈ Point(face_lattice(I)) 
  supposing a ≤ phi ∧ (∀[g:{f:I ⟶ I| (phi f) 1} ]. ((a)<g> u(g) ∈ Point(face_lattice(I))))

Lemma: extend-face-term-morph
[I:fset(ℕ)]. ∀[phi:Point(face_lattice(I))]. ∀[u:{I,phi ⊢ _:𝔽}]. ∀[J:fset(ℕ)]. ∀[f:J ⟶ I].
  ((extend-face-term(I;phi;u))<f> extend-face-term(J;(phi)<f>;(u)subset-trans(I;J;f;phi)) ∈ Point(face_lattice(J)))

Definition: face_comp
face_comp() ==  λI,i,rho,phi,u,a0. (i1)(extend-face-term(I+i;s(phi);u))

Lemma: face_comp_wf
[G:j⊢]. (face_comp() ∈ G ⊢ CompOp(𝔽))

Definition: face-comp
face-comp() ==  cop-to-cfun(face_comp())

Lemma: face-comp_wf
[G:j⊢]. (face-comp() ∈ G ⊢ Compositon(𝔽))

Definition: bij-contr
bij-contr(X; A; f; g; cA; b; c) ==
  contr-witness(X;app(g; contr-center(c));map-path(X.A;(g)p;contr-path((c)p;app((f)p; q))) b)

Lemma: bij-contr_wf
X:j⊢. ∀A,B:{X ⊢ _}. ∀f:{X ⊢ _:(A ⟶ B)}. ∀g:{X ⊢ _:(B ⟶ A)}. ∀cA:X +⊢ Compositon(A).
b:{X.A ⊢ _:(Path_(A)p app((g)p; app((f)p; q)) q)}. ∀c:{X ⊢ _:Contractible(B)}.
  (bij-contr(X; A; f; g; cA; b; c) ∈ {X ⊢ _:Contractible(A)})

Lemma: bijection-preserves-contractible
X:j⊢. ∀A,B:{X ⊢ _}. ∀f:{X ⊢ _:(A ⟶ B)}. ∀g:{X ⊢ _:(B ⟶ A)}. ∀cA:X +⊢ Compositon(A).
b:{X.A ⊢ _:(Path_(A)p app((g)p; app((f)p; q)) q)}. ∀c:{X ⊢ _:Contractible(B)}.
  {X ⊢ _:Contractible(A)}

Comment: pres doc
=================================  pres  ======================================
 The cubical term pres [phi ⊢→ t] t0  is witness to the fact that 
composition is `preserved by` function application.⋅

Definition: pres-a0
pres-a0(G;f;t0) ==  app((f)[0(𝕀)]; t0)

Lemma: pres-a0_wf
[G:j⊢]. ∀[A,T:{G.𝕀 ⊢ _}]. ∀[f:{G.𝕀 ⊢ _:(T ⟶ A)}]. ∀[t0:{G ⊢ _:(T)[0(𝕀)]}].  (pres-a0(G;f;t0) ∈ {G ⊢ _:(A)[0(𝕀)]})

Definition: pres-c1
pres-c1(G;phi;f;t;t0;cA) ==  comp cA [phi ⊢→ app(f; t)] app((f)[0(𝕀)]; t0)

Lemma: pres-c1_wf
[G:j⊢]. ∀[phi:{G ⊢ _:𝔽}]. ∀[A,T:{G.𝕀 ⊢ _}]. ∀[f:{G.𝕀 ⊢ _:(T ⟶ A)}]. ∀[t:{G.𝕀(phi)p ⊢ _:T}].
[t0:{G ⊢ _:(T)[0(𝕀)][phi |⟶ t[0]]}]. ∀[cA:composition-function{j:l,i:l}(G.𝕀;A)].
  (pres-c1(G;phi;f;t;t0;cA) ∈ {G ⊢ _:(A)[1(𝕀)][phi |⟶ app(f; t)[1]]})

Lemma: csm-pres-c1
[G:j⊢]. ∀[phi:{G ⊢ _:𝔽}]. ∀[A,T:{G.𝕀 ⊢ _}]. ∀[f:{G.𝕀 ⊢ _:(T ⟶ A)}]. ∀[t:{G.𝕀(phi)p ⊢ _:T}].
[t0:{G ⊢ _:(T)[0(𝕀)][phi |⟶ t[0]]}]. ∀[cA:G.𝕀 ⊢ Compositon(A)]. ∀[H:j⊢]. ∀[s:H j⟶ G].
  ((pres-c1(G;phi;f;t;t0;cA))s
  pres-c1(H;(phi)s;(f)s+;(t)s+;(t0)s;(cA)s+)
  ∈ {H ⊢ _:((A)s+)[1(𝕀)][(phi)s |⟶ app((f)s+; (t)s+)[1]]})

Definition: pres-c2
pres-c2(G;phi;f;t;t0;cT) ==  app((f)[1(𝕀)]; comp cT [phi ⊢→ t] t0)

Lemma: pres-c2_wf
[G:j⊢]. ∀[phi:{G ⊢ _:𝔽}]. ∀[A,T:{G.𝕀 ⊢ _}]. ∀[f:{G.𝕀 ⊢ _:(T ⟶ A)}]. ∀[t:{G.𝕀(phi)p ⊢ _:T}].
[t0:{G ⊢ _:(T)[0(𝕀)][phi |⟶ t[0]]}]. ∀[cT:composition-function{j:l,i:l}(G.𝕀;T)].
  (pres-c2(G;phi;f;t;t0;cT) ∈ {G ⊢ _:(A)[1(𝕀)][phi |⟶ app(f; t)[1]]})

Lemma: csm-pres-c2
[G:j⊢]. ∀[phi:{G ⊢ _:𝔽}]. ∀[A,T:{G.𝕀 ⊢ _}]. ∀[f:{G.𝕀 ⊢ _:(T ⟶ A)}]. ∀[t:{G.𝕀(phi)p ⊢ _:T}].
[t0:{G ⊢ _:(T)[0(𝕀)][phi |⟶ t[0]]}]. ∀[cT:G.𝕀 ⊢ Compositon(T)]. ∀[H:j⊢]. ∀[s:H j⟶ G].
  ((pres-c2(G;phi;f;t;t0;cT))s
  pres-c2(H;(phi)s;(f)s+;(t)s+;(t0)s;(cT)s+)
  ∈ {H ⊢ _:((A)s+)[1(𝕀)][(phi)s |⟶ app((f)s+; (t)s+)[1]]})

Definition: pres-v
pres-v(G;phi;t;t0;cT) ==  fill cT [phi ⊢→ t] t0

Lemma: pres-v_wf
[G:j⊢]. ∀[phi:{G ⊢ _:𝔽}]. ∀[T:{G.𝕀 ⊢ _}]. ∀[t:{G.𝕀(phi)p ⊢ _:T}]. ∀[t0:{G ⊢ _:(T)[0(𝕀)][phi |⟶ t[0]]}].
[cT:composition-function{j:l,i:l}(G.𝕀;T)].
  (pres-v(G;phi;t;t0;cT) ∈ {G.𝕀 ⊢ _:T[(phi)p |⟶ t]})

Lemma: csm-pres-v
[G:j⊢]. ∀[phi:{G ⊢ _:𝔽}]. ∀[T:{G.𝕀 ⊢ _}]. ∀[t:{G.𝕀(phi)p ⊢ _:T}]. ∀[t0:{G ⊢ _:(T)[0(𝕀)][phi |⟶ t[0]]}].
[cT:G.𝕀 ⊢ Compositon(T)]. ∀[H:j⊢]. ∀[s:H j⟶ G].
  ((pres-v(G;phi;t;t0;cT))s+ pres-v(H;(phi)s;(t)s+;(t0)s;(cT)s+) ∈ {H.𝕀 ⊢ _:(T)s+[((phi)s)p |⟶ (t)s+]})

Definition: presw
presw(G;phi;f;t;t0;cT) ==  app(f; pres-v(G;phi;t;t0;cT))

Lemma: presw_wf
[G:j⊢]. ∀[phi:{G ⊢ _:𝔽}]. ∀[A,T:{G.𝕀 ⊢ _}]. ∀[f:{G.𝕀 ⊢ _:(T ⟶ A)}]. ∀[t:{G.𝕀(phi)p ⊢ _:T}].
[t0:{G ⊢ _:(T)[0(𝕀)][phi |⟶ t[0]]}]. ∀[cT:composition-function{j:l,i:l}(G.𝕀;T)].
  (presw(G;phi;f;t;t0;cT) ∈ {G.𝕀 ⊢ _:A})

Lemma: csm-presw
[G:j⊢]. ∀[phi:{G ⊢ _:𝔽}]. ∀[A,T:{G.𝕀 ⊢ _}]. ∀[f:{G.𝕀 ⊢ _:(T ⟶ A)}]. ∀[t:{G.𝕀(phi)p ⊢ _:T}].
[t0:{G ⊢ _:(T)[0(𝕀)][phi |⟶ t[0]]}]. ∀[cT:G.𝕀 ⊢ Compositon(T)]. ∀[H:j⊢]. ∀[s:H j⟶ G].
  ((presw(G;phi;f;t;t0;cT))s+ presw(H;(phi)s;(f)s+;(t)s+;(t0)s;(cT)s+) ∈ {H.𝕀 ⊢ _:(A)s+})

Lemma: pres-a0-constraint
[G:j⊢]. ∀[phi:{G ⊢ _:𝔽}]. ∀[A,T:{G.𝕀 ⊢ _}]. ∀[f:{G.𝕀 ⊢ _:(T ⟶ A)}]. ∀[t:{G.𝕀(phi)p ⊢ _:T}].
[t0:{G ⊢ _:(T)[0(𝕀)][phi |⟶ t[0]]}]. ∀[cT:G.𝕀 +⊢ Compositon(T)].
  ((pres-a0(G;f;t0))p ∈ {G.𝕀 ⊢ _:((A)p+)[0(𝕀)][((phi)p ∨ (q=1)) |⟶ ((presw(G;phi;f;t;t0;cT))p+)[0(𝕀)]]})

Lemma: presw-pres-c2
[G:j⊢]. ∀[phi:{G ⊢ _:𝔽}]. ∀[A,T:{G.𝕀 ⊢ _}]. ∀[f:{G.𝕀 ⊢ _:(T ⟶ A)}]. ∀[t:{G.𝕀(phi)p ⊢ _:T}].
[t0:{G ⊢ _:(T)[0(𝕀)][phi |⟶ t[0]]}]. ∀[cT:G.𝕀 ⊢ Compositon(T)].
  ((((presw(G;phi;f;t;t0;cT))p+)[1(𝕀)])[1(𝕀)] pres-c2(G;phi;f;t;t0;cT) ∈ {G ⊢ _:(A)[1(𝕀)]})

Lemma: presw-pres-c1
[G:j⊢]. ∀[phi:{G ⊢ _:𝔽}]. ∀[A,T:{G.𝕀 ⊢ _}]. ∀[f:{G.𝕀 ⊢ _:(T ⟶ A)}]. ∀[t:{G.𝕀(phi)p ⊢ _:T}].
[t0:{G ⊢ _:(T)[0(𝕀)][phi |⟶ t[0]]}]. ∀[cT:G.𝕀 +⊢ Compositon(T)]. ∀[cA:G.𝕀 ⊢ Compositon(A)].
  G ⊢ (comp (cA)p+ [((phi)p ∨ (q=1)) ⊢→ (presw(G;phi;f;t;t0;cT))p+] (pres-a0(G;f;t0))p)[0(𝕀)]=pres-c1(G;phi;f;t;t0;cA):
  (A)[1(𝕀)]

Definition: pres
pres [phi ⊢→ t] t0 ==  <>(comp (cA)p+ [((phi)p ∨ (q=1)) ⊢→ (presw(G;phi;f;t;t0;cT))p+] (pres-a0(G;f;t0))p)

Lemma: pres_wf
[G:j⊢]. ∀[phi:{G ⊢ _:𝔽}]. ∀[A,T:{G.𝕀 ⊢ _}]. ∀[f:{G.𝕀 ⊢ _:(T ⟶ A)}]. ∀[t:{G.𝕀(phi)p ⊢ _:T}].
[t0:{G ⊢ _:(T)[0(𝕀)][phi |⟶ t[0]]}]. ∀[cT:G.𝕀 +⊢ Compositon(T)]. ∀[cA:G.𝕀 +⊢ Compositon(A)].
  (pres [phi ⊢→ t] t0 ∈ {G ⊢ _:(Path_(A)[1(𝕀)] pres-c1(G;phi;f;t;t0;cA) pres-c2(G;phi;f;t;t0;cT))})

Lemma: csm-pres
[G:j⊢]. ∀[phi:{G ⊢ _:𝔽}]. ∀[A,T:{G.𝕀 ⊢ _}]. ∀[f:{G.𝕀 ⊢ _:(T ⟶ A)}]. ∀[t:{G.𝕀(phi)p ⊢ _:T}].
[t0:{G ⊢ _:(T)[0(𝕀)][phi |⟶ t[0]]}]. ∀[cT:G.𝕀 +⊢ Compositon(T)]. ∀[cA:G.𝕀 +⊢ Compositon(A)]. ∀[H:j⊢]. ∀[s:H j⟶ G].
  ((pres [phi ⊢→ t] t0)s
  pres (f)s+ [(phi)s ⊢→ (t)s+] (t0)s
  ∈ {H ⊢ _:(Path_((A)s+)[1(𝕀)] pres-c1(H;(phi)s;(f)s+;(t)s+;(t0)s;(cA)s+) pres-c2(H;(phi)s;(f)s+;(t)s+;(t0)s;(cT)s+))})

Lemma: pres-invariant
[G,H:j⊢].
  ∀[phi:{G ⊢ _:𝔽}]. ∀[A,T:{G.𝕀 ⊢ _}]. ∀[f:{G.𝕀 ⊢ _:(T ⟶ A)}]. ∀[t:{G.𝕀(phi)p ⊢ _:T}].
  ∀[t0:{G ⊢ _:(T)[0(𝕀)][phi |⟶ t[0]]}]. ∀[cT:G.𝕀 +⊢ Compositon(T)]. ∀[cA:G.𝕀 +⊢ Compositon(A)].
    (pres [phi ⊢→ t] t0
    pres [phi ⊢→ t] t0
    ∈ {G ⊢ _:(Path_(A)[1(𝕀)] pres-c1(G;phi;f;t;t0;cA) pres-c2(G;phi;f;t;t0;cT))}) 
  supposing G ∈ CubicalSet{j}

Lemma: pres-constraint
[G:j⊢]. ∀[phi:{G ⊢ _:𝔽}]. ∀[A,T:{G.𝕀 ⊢ _}]. ∀[f:{G.𝕀 ⊢ _:(T ⟶ A)}]. ∀[t:{G.𝕀(phi)p ⊢ _:T}].
[t0:{G ⊢ _:(T)[0(𝕀)][phi |⟶ t[0]]}]. ∀[cT:G.𝕀 +⊢ Compositon(T)]. ∀[cA:G.𝕀 +⊢ Compositon(A)].
  (pres [phi ⊢→ t] t0
  G ⊢ <>((app(f; t)[1])p)
  ∈ {G, phi ⊢ _:(Path_(A)[1(𝕀)] pres-c1(G;phi;f;t;t0;cA) pres-c2(G;phi;f;t;t0;cT))})

Lemma: pres_wf2
[G:j⊢]. ∀[phi:{G ⊢ _:𝔽}]. ∀[A,T:{G.𝕀 ⊢ _}]. ∀[f:{G.𝕀 ⊢ _:(T ⟶ A)}]. ∀[t:{G.𝕀(phi)p ⊢ _:T}].
[t0:{G ⊢ _:(T)[0(𝕀)][phi |⟶ t[0]]}]. ∀[cT:G.𝕀 +⊢ Compositon(T)]. ∀[cA:G.𝕀 +⊢ Compositon(A)].
  (pres [phi ⊢→ t] t0 ∈ {G ⊢ _:(Path_(A)[1(𝕀)] pres-c1(G;phi;f;t;t0;cA) pres-c2(G;phi;f;t;t0;cT))[phi 
                                 |⟶ <>((app(f; t)[1])p)]})

Definition: equiv-term
equiv [phi ⊢→ (t,  c)] ==
  let equiv-contr(f;a) in
   let (contr-path(e;fiber-point(t;c)))p in
   comp (cF)p [phi ⊢→ u] contr-center(e)

Lemma: equiv-term_wf
[G:j⊢]. ∀[phi:{G ⊢ _:𝔽}]. ∀[A,T:{G ⊢ _}]. ∀[f:{G ⊢ _:Equiv(T;A)}]. ∀[t:{G, phi ⊢ _:T}]. ∀[a:{G ⊢ _:A}].
[c:{G, phi ⊢ _:(Path_A app(equiv-fun(f); t))}]. ∀[cF:G ⊢ Compositon(Fiber(equiv-fun(f);a))].
  (equiv [phi ⊢→ (t,  c)] a ∈ {G ⊢ _:Fiber(equiv-fun(f);a)[phi |⟶ fiber-point(t;c)]})

Lemma: equiv-term-subset
[G:j⊢]. ∀[phi:{G ⊢ _:𝔽}]. ∀[A,T:{G ⊢ _}]. ∀[f:{G ⊢ _:Equiv(T;A)}]. ∀[t:{G, phi ⊢ _:T}]. ∀[a:{G ⊢ _:A}].
[c:{G, phi ⊢ _:(Path_A app(equiv-fun(f); t))}]. ∀[cF:G ⊢ Compositon(Fiber(equiv-fun(f);a))]. ∀[psi:{G ⊢ _:𝔽}].
  (equiv [phi ⊢→ (t,  c)] equiv [phi ⊢→ (t,  c)] a ∈ {G, psi ⊢ _:Fiber(equiv-fun(f);a)})

Lemma: equiv-term-0
[G:j⊢]. ∀[phi:{G ⊢ _:𝔽}].
  ∀[A,T:{G ⊢ _}]. ∀[f:{G ⊢ _:Equiv(T;A)}]. ∀[a:{G ⊢ _:A}]. ∀[t,c:Top]. ∀[cF:G +⊢ Compositon(Fiber(equiv-fun(f);a))].
    (equiv [phi ⊢→ (t,  c)] transprt(G;(cF)p;contr-center(equiv-contr(f;a))) ∈ {G ⊢ _:Fiber(equiv-fun(f);a)}) 
  supposing phi 0(𝔽) ∈ {G ⊢ _:𝔽}

Lemma: equiv-term-0-subset-1
[G:j⊢]. ∀[phi:{G ⊢ _:𝔽}].
  ∀[psi:{G ⊢ _:𝔽}]
    ∀[A,T:{G ⊢ _}]. ∀[f:{G ⊢ _:Equiv(T;A)}]. ∀[a:{G ⊢ _:A}]. ∀[t,c:Top]. ∀[cF:G +⊢ Compositon(Fiber(equiv-fun(f);a))].
      (equiv [phi ⊢→ (t,  c)] transprt(G;(cF)p;contr-center(equiv-contr(f;a))) ∈ {G ⊢ _:Fiber(equiv-fun(f);a)}) 
    supposing psi 1(𝔽) ∈ {G ⊢ _:𝔽
  supposing phi 0(𝔽) ∈ {G ⊢ _:𝔽}

Definition: equiv_term
equiv [phi ⊢→ (t,c)] ==  equiv [phi ⊢→ (t,  c)] a

Lemma: equiv_term_wf
[G:j⊢]. ∀[phi:{G ⊢ _:𝔽}]. ∀[A,T:{G ⊢ _}]. ∀[f:{G ⊢ _:Equiv(T;A)}]. ∀[t:{G, phi ⊢ _:T}]. ∀[a:{G ⊢ _:A}].
[c:{G, phi ⊢ _:(Path_A app(equiv-fun(f); t))}]. ∀[cA:G +⊢ Compositon(A)]. ∀[cT:G +⊢ Compositon(T)].
  (equiv [phi ⊢→ (t,c)] a ∈ {G ⊢ _:Fiber(equiv-fun(f);a)[phi |⟶ fiber-point(t;c)]})

Lemma: equiv_term-0
[G:j⊢]. ∀[phi:{G ⊢ _:𝔽}].
  ∀[A,T:{G ⊢ _}]. ∀[f:{G ⊢ _:Equiv(T;A)}]. ∀[a:{G ⊢ _:A}]. ∀[t,c:Top]. ∀[cA:G +⊢ Compositon(A)].
  ∀[cT:G +⊢ Compositon(T)].
    (equiv [phi ⊢→ (t,c)] a
    transprt(G;(fiber-comp(G;T;A;equiv-fun(f);a;cT;cA))p;contr-center(equiv-contr(f;a)))
    ∈ {G ⊢ _:Fiber(equiv-fun(f);a)}) 
  supposing phi 0(𝔽) ∈ {G ⊢ _:𝔽}

Definition: csm-path-ap-q
csm-path-ap-q(H;G;s;t) ==  ((t)p)s+ q

Lemma: csm-path-ap-q_wf
[G:j⊢]. ∀[phi:{G ⊢ _:𝔽}]. ∀[H:j⊢]. ∀[s:H j⟶ G]. ∀[B:{G ⊢ _}]. ∀[a,b:{G, phi ⊢ _:B}]. ∀[t:{G, phi ⊢ _:(Path_B b)}].
  (csm-path-ap-q(H;G;s;t) ∈ {H.𝕀((phi)p)s+ ⊢ _:((B)p)s+})

Lemma: csm-equiv-term
[G:j⊢]. ∀[phi:{G ⊢ _:𝔽}]. ∀[A,T:{G ⊢ _}]. ∀[f:{G ⊢ _:Equiv(T;A)}]. ∀[t:{G, phi ⊢ _:T}]. ∀[a:{G ⊢ _:A}].
[c:{G, phi ⊢ _:(Path_A app(equiv-fun(f); t))}]. ∀[cF:G ⊢ Compositon(Fiber(equiv-fun(f);a))]. ∀[H:j⊢]. ∀[s:H j⟶ G].
  ((equiv [phi ⊢→ (t,  c)] a)s
  equiv (f)s [(phi)s ⊢→ ((t)s,  (c)s)] (a)s
  ∈ {H ⊢ _:Fiber(equiv-fun((f)s);(a)s)[(phi)s |⟶ fiber-point((t)s;(c)s)]})

Lemma: csm-equiv_term
[G:j⊢]. ∀[phi:{G ⊢ _:𝔽}]. ∀[A,T:{G ⊢ _}]. ∀[f:{G ⊢ _:Equiv(T;A)}]. ∀[t:{G, phi ⊢ _:T}]. ∀[a:{G ⊢ _:A}].
[c:{G, phi ⊢ _:(Path_A app(equiv-fun(f); t))}]. ∀[cA:G +⊢ Compositon(A)]. ∀[cT:G +⊢ Compositon(T)]. ∀[H:j⊢].
[s:H j⟶ G].
  ((equiv [phi ⊢→ (t,c)] a)s
  equiv (f)s [(phi)s ⊢→ ((t)s,(c)s)] (a)s
  ∈ {H ⊢ _:Fiber(equiv-fun((f)s);(a)s)[(phi)s |⟶ fiber-point((t)s;(c)s)]})

Comment: glueing doc
=================================  glueing  ===================================

The hardest theorems of the theory are in this section where we define the
cubical type Glue [phi ⊢→ (T;w)] and show that it has composition structure.



Definition: glue-equations
glue-equations(Gamma;A;phi;T;w;I;rho;t;a) ==
  (∀J:fset(ℕ). ∀f:{f:J ⟶ I| phi(f(rho)) 1 ∈ Point(face_lattice(J))} . ∀K:fset(ℕ). ∀g:K ⟶ J.
     ((t f(rho) g) (t f ⋅ g) ∈ T(f ⋅ g(rho))))
  ∧ (∀J:fset(ℕ). ∀f:{f:J ⟶ I| phi(f(rho)) 1 ∈ Point(face_lattice(J))} .
       ((w(f(rho)) (t f)) (a rho f) ∈ A(f(rho))))

Lemma: glue-equations_wf
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[phi:{Gamma ⊢ _:𝔽}]. ∀[T:{Gamma, phi ⊢ _}]. ∀[w:{Gamma, phi ⊢ _:(T ⟶ A)}].
[I:fset(ℕ)]. ∀[rho:Gamma(I)]. ∀[t:J:fset(ℕ) ⟶ f:{f:J ⟶ I| phi(f(rho)) 1 ∈ Point(face_lattice(J))}  ⟶ T(f(rho))].
[a:A(rho)].
  (glue-equations(Gamma;A;phi;T;w;I;rho;t;a) ∈ ℙ)

Definition: glue-cube
glue-cube(Gamma;A;phi;T;w;I;rho) ==
  if (phi(rho)==1)
  then T(rho)
  else {ta:J:fset(ℕ) ⟶ f:{f:J ⟶ I| phi(f(rho)) 1 ∈ Point(face_lattice(J))}  ⟶ T(f(rho)) × A(rho)| 
        let t,a ta 
        in glue-equations(Gamma;A;phi;T;w;I;rho;t;a)} 
  fi 

Lemma: glue-cube_wf
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[phi:{Gamma ⊢ _:𝔽}]. ∀[T:{Gamma, phi ⊢ _}]. ∀[w:{Gamma, phi ⊢ _:(T ⟶ A)}].
[I:fset(ℕ)]. ∀[rho:Gamma(I)].
  (glue-cube(Gamma;A;phi;T;w;I;rho) ∈ Type)

Lemma: equal-glue-cube
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}].
  ∀phi:{Gamma ⊢ _:𝔽}
    ∀[T:{Gamma, phi ⊢ _}]. ∀[w:{Gamma, phi ⊢ _:(T ⟶ A)}].
      ∀I:fset(ℕ). ∀rho:Gamma(I). ∀u,v:glue-cube(Gamma;A;phi;T;w;I;rho).
        v ∈ glue-cube(Gamma;A;phi;T;w;I;rho) 
        supposing if (phi(rho)==1)
        then v ∈ T(rho)
        else v ∈ (J:fset(ℕ) ⟶ f:{f:J ⟶ I| phi(f(rho)) 1 ∈ Point(face_lattice(J))}  ⟶ T(f(rho)) × A(rho))
        fi 

Definition: glue-morph
glue-morph(Gamma;A;phi;T;w;I;rho;J;f;u) ==
  if (phi(rho)==1)
  then (u rho f)
  else let t,a 
       in if (phi(f(rho))==1) then else <λK,g. (t f ⋅ g), (a rho f)> fi 
  fi 

Lemma: glue-morph_wf
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[phi:{Gamma ⊢ _:𝔽}]. ∀[T:{Gamma, phi ⊢ _}]. ∀[w:{Gamma, phi ⊢ _:(T ⟶ A)}].
[I:fset(ℕ)]. ∀[rho:Gamma(I)]. ∀[J:fset(ℕ)]. ∀[f:J ⟶ I]. ∀[u:glue-cube(Gamma;A;phi;T;w;I;rho)].
  (glue-morph(Gamma;A;phi;T;w;I;rho;J;f;u) ∈ glue-cube(Gamma;A;phi;T;w;J;f(rho)))

Lemma: glue-morph-comp
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[phi:{Gamma ⊢ _:𝔽}]. ∀[T:{Gamma, phi ⊢ _}]. ∀[w:{Gamma, phi ⊢ _:(T ⟶ A)}].
[I:fset(ℕ)]. ∀[rho:Gamma(I)]. ∀[J,K:fset(ℕ)]. ∀[f:J ⟶ I]. ∀[g:K ⟶ J]. ∀[u:glue-cube(Gamma;A;phi;T;w;I;rho)].
  (glue-morph(Gamma;A;phi;T;w;J;f(rho);K;g;glue-morph(Gamma;A;phi;T;w;I;rho;J;f;u))
  glue-morph(Gamma;A;phi;T;w;I;rho;K;f ⋅ g;u)
  ∈ glue-cube(Gamma;A;phi;T;w;K;f ⋅ g(rho)))

Lemma: glue-morph-id
Gamma:j⊢. ∀A:{Gamma ⊢ _}. ∀phi:{Gamma ⊢ _:𝔽}. ∀T:{Gamma, phi ⊢ _}. ∀w:{Gamma, phi ⊢ _:(T ⟶ A)}. ∀I:fset(ℕ).
a:Gamma(I). ∀u:glue-cube(Gamma;A;phi;T;w;I;a).
  (glue-morph(Gamma;A;phi;T;w;I;a;I;1;u) u ∈ glue-cube(Gamma;A;phi;T;w;I;a))

Definition: glue-type
Glue [phi ⊢→ (T;w)] ==  <λI,rho. glue-cube(Gamma;A;phi;T;w;I;rho), λI,J,f,rho,u. glue-morph(Gamma;A;phi;T;w;I;rho;J;f;\000Cu)>

Lemma: glue-type_wf
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[phi:{Gamma ⊢ _:𝔽}]. ∀[T:{Gamma, phi ⊢ _}]. ∀[w:{Gamma, phi ⊢ _:(T ⟶ A)}].
  Gamma ⊢ Glue [phi ⊢→ (T;w)] A

Lemma: glue-type-subset
[Gamma,A,phi,T,w,xx:Top].  (Gamma, xx ⊢ Glue [phi ⊢→ (T;w)] Gamma ⊢ Glue [phi ⊢→ (T;w)] A)

Lemma: csm-glue-type
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[phi:{Gamma ⊢ _:𝔽}]. ∀[T:{Gamma, phi ⊢ _}]. ∀[w:{Gamma, phi ⊢ _:(T ⟶ A)}]. ∀[Z:j⊢].
[s:Z j⟶ Gamma].
  ((Glue [phi ⊢→ (T;w)] A)s Z ⊢ Glue [(phi)s ⊢→ ((T)s;(w)s)] (A)s ∈ {Z ⊢ _})

Lemma: glue-type-constraint
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[phi:{Gamma ⊢ _:𝔽}]. ∀[T:{Gamma, phi ⊢ _}]. ∀[w:{Gamma, phi ⊢ _:(T ⟶ A)}].
  Gamma, phi ⊢ Glue [phi ⊢→ (T;w)] T

Lemma: glue-type-term-subtype
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[phi:{Gamma ⊢ _:𝔽}]. ∀[T:{Gamma, phi ⊢ _}]. ∀[w:{Gamma, phi ⊢ _:(T ⟶ A)}].
  ({Gamma ⊢ _:Glue [phi ⊢→ (T;w)] A} ⊆{Gamma, phi ⊢ _:T})

Lemma: glue-type-term-subtype2
[Gamma:j⊢]. ∀[psi:{Gamma ⊢ _:𝔽}]. ∀[A:{Gamma, psi ⊢ _}]. ∀[phi:{Gamma, psi ⊢ _:𝔽}]. ∀[T:{Gamma, psi, phi ⊢ _}].
[w:{Gamma, psi, phi ⊢ _:(T ⟶ A)}].
  ({Gamma, psi ⊢ _:Glue [phi ⊢→ (T;w)] A} ⊆{Gamma, psi, phi ⊢ _:T})

Lemma: glue-type-1
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[T:{Gamma, 1(𝔽) ⊢ _}]. ∀[w:{Gamma, 1(𝔽) ⊢ _:(T ⟶ A)}].
  (Gamma ⊢ Glue [1(𝔽) ⊢→ (T;w)] T ∈ {Gamma ⊢ _})

Lemma: glue-type-1'
[Gamma:j⊢]. ∀[A,T:{Gamma ⊢ _}]. ∀[w:{Gamma ⊢ _:(T ⟶ A)}]. ∀[phi:{Gamma ⊢ _:𝔽}].
  Gamma ⊢ Glue [phi ⊢→ (T;w)] T ∈ {Gamma ⊢ _} supposing phi 1(𝔽) ∈ {Gamma ⊢ _:𝔽}

Definition: glue-term
glue [phi ⊢→ t] ==  λI,rho. if (phi(rho)==1) then t(rho) else <λJ,f. t(f(rho)), a(rho)> fi 

Lemma: glue-term_wf
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[phi:{Gamma ⊢ _:𝔽}]. ∀[T:{Gamma, phi ⊢ _}]. ∀[w:{Gamma, phi ⊢ _:(T ⟶ A)}].
[t:{Gamma, phi ⊢ _:T}]. ∀[a:{Gamma ⊢ _:A[phi |⟶ app(w; t)]}].
  (glue [phi ⊢→ t] a ∈ {Gamma ⊢ _:Glue [phi ⊢→ (T;w)] A})

Lemma: csm-glue-term
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[phi:{Gamma ⊢ _:𝔽}]. ∀[T:{Gamma, phi ⊢ _}]. ∀[w:{Gamma, phi ⊢ _:(T ⟶ A)}].
[t:{Gamma, phi ⊢ _:T}]. ∀[a:{Gamma ⊢ _:A[phi |⟶ app(w; t)]}]. ∀[H:j⊢]. ∀[s:H j⟶ Gamma].
  ((glue [phi ⊢→ t] a)s H ⊢ glue [(phi)s ⊢→ (t)s] (a)s ∈ {H ⊢ _:(Glue [phi ⊢→ (T;w)] A)s})

Lemma: glue-term-constraint
[Gamma:j⊢]. ∀[phi:{Gamma ⊢ _:𝔽}]. ∀[T:{Gamma, phi ⊢ _}]. ∀[t:{Gamma, phi ⊢ _:T}]. ∀[A,a:Top].
  Gamma, phi ⊢ glue [phi ⊢→ t] a=t:T

Lemma: glue-term-subset
[Gamma:j⊢]. ∀[A,phi,T,t,a,xx:Top].  (Gamma, xx ⊢ glue [phi ⊢→ t] Gamma ⊢ glue [phi ⊢→ t] a)

Lemma: glue-term-1
[Gamma:j⊢]. ∀[T:{Gamma ⊢ _}]. ∀[t:{Gamma ⊢ _:T}]. ∀[A,a:Top].  (Gamma ⊢ glue [1(𝔽) ⊢→ t] t ∈ {Gamma ⊢ _:T})

Lemma: glue-term-1'
[Gamma:j⊢]. ∀[phi:{Gamma ⊢ _:𝔽}]. ∀[T:{Gamma ⊢ _}]. ∀[t:{Gamma ⊢ _:T}]. ∀[A,a:Top].
  Gamma ⊢ glue [phi ⊢→ t] t ∈ {Gamma ⊢ _:T} supposing phi 1(𝔽) ∈ {Gamma ⊢ _:𝔽}

Lemma: glue-term_wf2
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[phi:{Gamma ⊢ _:𝔽}]. ∀[T:{Gamma, phi ⊢ _}]. ∀[w:{Gamma, phi ⊢ _:(T ⟶ A)}].
[t:{Gamma, phi ⊢ _:T}]. ∀[a:{Gamma ⊢ _:A[phi |⟶ app(w; t)]}].
  (glue [phi ⊢→ t] a ∈ {Gamma ⊢ _:Glue [phi ⊢→ (T;w)] A[phi |⟶ t]})

Definition: unglue-term
unglue(b) ==  λI,rho. if (phi(rho)==1) then rho b(rho) else snd(b(rho)) fi 

Lemma: unglue-term_wf
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[phi:{Gamma ⊢ _:𝔽}]. ∀[T:{Gamma, phi ⊢ _}]. ∀[w:{Gamma, phi ⊢ _:(T ⟶ A)}].
[b:{Gamma ⊢ _:Glue [phi ⊢→ (T;w)] A}].
  (unglue(b) ∈ {Gamma ⊢ _:A})

Lemma: unglue-term_wf2
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[phi:{Gamma ⊢ _:𝔽}]. ∀[T:{Gamma, phi ⊢ _}]. ∀[w:{Gamma, phi ⊢ _:(T ⟶ A)}].
[b:{Gamma ⊢ _:Glue [phi ⊢→ (T;w)] A}].
  (unglue(b) ∈ {Gamma ⊢ _:A[phi |⟶ app(w; b)]})

Lemma: unglue-term-1
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[phi:{Gamma ⊢ _:𝔽}]. ∀[T:{Gamma, phi ⊢ _}]. ∀[w:{Gamma, phi ⊢ _:(T ⟶ A)}].
[b:{Gamma ⊢ _:Glue [phi ⊢→ (T;w)] A}].
  unglue(b) app(w; b) ∈ {Gamma ⊢ _:A} supposing phi 1(𝔽) ∈ {Gamma ⊢ _:𝔽}

Lemma: csm-unglue
[phi,w,s,b:Top].  ((unglue(b))s unglue((b)s))

Lemma: unglue-glue
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[phi:{Gamma ⊢ _:𝔽}]. ∀[T:{Gamma, phi ⊢ _}]. ∀[w:{Gamma, phi ⊢ _:(T ⟶ A)}].
[t:{Gamma, phi ⊢ _:T}]. ∀[a:{Gamma ⊢ _:A[phi |⟶ app(w; t)]}].
  (unglue(glue [phi ⊢→ t] a) a ∈ {Gamma ⊢ _:A})

Lemma: glue-unglue
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[phi:{Gamma ⊢ _:𝔽}]. ∀[T:{Gamma, phi ⊢ _}]. ∀[w:{Gamma, phi ⊢ _:(T ⟶ A)}].
[b:{Gamma ⊢ _:Glue [phi ⊢→ (T;w)] A}].
  (Gamma ⊢ glue [phi ⊢→ b] unglue(b) b ∈ {Gamma ⊢ _:Glue [phi ⊢→ (T;w)] A})

Definition: glue-comp
comp(Glue [phi ⊢→ (T, f)] A)  ==
  λH,sigma,psi,b,b0. let unglue(b) in
                      let a0 unglue(b0) in
                      let a'1 comp (cA)sigma [psi ⊢→ a] a0 in
                      let t'1 comp (cT)sigma [psi ⊢→ b] b0 in
                      let pres (equiv-fun(f))sigma [psi ⊢→ b] b0 in
                      let st (t'1 ∨ (b)[1(𝕀)]) in
                      let sw (w ∨ <>((app((equiv-fun(f))sigma; b)[1])p)) in
                      let cF fiber-comp(H, ((phi)sigma)[1(𝕀)];((T)sigma)[1(𝕀)];((A)sigma)[1(𝕀)];
                                          equiv-fun(((f)sigma)[1(𝕀)]);a'1;(cT)sigma [1(𝕀)];(cA)sigma [1(𝕀)]) in
                      let equiv ((f)sigma)[1(𝕀)] [((∀ (phi)sigma) ∨ psi) ⊢→ (st,  sw)] a'1 in
                      let t1 fiber-member(z) in
                      let alpha fiber-path(z) in
                      let a1 comp (cA)sigma [1(𝕀)] [(((phi)sigma)[1(𝕀)] ∨ psi) ⊢→ ((alpha)p q ∨ (a[1])p)]
                                   a'1 in
                      glue [((phi)sigma)[1(𝕀)] ⊢→ t1] a1

Lemma: glue-comp_wf
G:j⊢. ∀A:{G ⊢ _}. ∀cA:G +⊢ Compositon(A). ∀psi:{G ⊢ _:𝔽}. ∀T:{G, psi ⊢ _}. ∀cT:G, psi +⊢ Compositon(T).
f:{G, psi ⊢ _:Equiv(T;A)}.
  (comp(Glue [psi ⊢→ (T, f)] A)  ∈ composition-function{j:l,i:l}(G;Glue [psi ⊢→ (T;equiv-fun(f))] A))

Comment: glue-comp-wf2-step2-defn
let  addr 121121222121311213211112 in let step2 make_sub_prf
   
(make_proof_address_term_aux (ioid Error :glue-comp_wf2) addr) in
 (view_show ioid_term) step2

Comment: glue-comp-wf2-step3-def
let  addr 221121222111211222 in (view_show ioid_term)(make_sub_prf  
(make_proof_address_term_aux (ioid Error :glue-comp_wf2-step2) addr))  


Comment: glue-comp-wf2-step4-def
let  addr 211722222121 in (view_show ioid_term)(make_sub_prf  
(make_proof_address_term_aux (ioid Error :glue-comp_wf2-step3) addr))  


Lemma: glue-comp_wf2
G:j⊢. ∀A:{G ⊢ _}. ∀cA:G +⊢ Compositon(A). ∀psi:{G ⊢ _:𝔽}. ∀T:{G, psi ⊢ _}. ∀cT:G, psi +⊢ Compositon(T).
f:{G, psi ⊢ _:Equiv(T;A)}.
  (comp(Glue [psi ⊢→ (T, f)] A)  ∈ G ⊢ Compositon(Glue [psi ⊢→ (T;equiv-fun(f))] A))

Lemma: glue-comp-agrees
The type (not displayed) of the equality in this lemma 
is composition-structure{i:l}(G, psi; T)
This means that in "extent" psi, when ⌜G ⊢ Glue [psi ⊢→ (T;f)] T ∈ {G, psi ⊢ _}⌝the
composition for ⌜Glue [psi ⊢→ (T;f)] A⌝ is the same as the composition for T.

This property of the compostion for Glue is used in construction of the 
composition for c𝕌  (the cubiucal universe type).⋅

[G:j⊢]. ∀[A:{G ⊢ _}]. ∀[cA:G +⊢ Compositon(A)]. ∀[psi:{G ⊢ _:𝔽}]. ∀[T:{G, psi ⊢ _}]. ∀[cT:G, psi +⊢ Compositon(T)].
[f:{G, psi ⊢ _:Equiv(T;A)}].
  (comp(Glue [psi ⊢→ (T, f)] A)  cT ∈ G, psi ⊢ Compositon(T))

Lemma: glue-comp-agrees2
The type (not displayed) of the equality in this lemma 
is composition-structure{i:l}(G, psi; T)
This means that in "extent" psi, when ⌜G ⊢ Glue [psi ⊢→ (T;f)] T ∈ {G, psi ⊢ _}⌝the
composition for ⌜Glue [psi ⊢→ (T;f)] A⌝ is the same as the composition for T.

This property of the compostion for Glue is used in construction of the 
composition for c𝕌  (the cubiucal universe type).⋅

[G:j⊢]. ∀[A:{G ⊢ _}]. ∀[cA:G +⊢ Compositon(A)]. ∀[psi:{G ⊢ _:𝔽}]. ∀[T:{G, psi ⊢ _}]. ∀[cT:G, psi +⊢ Compositon(T)].
[f:{G, psi ⊢ _:Equiv(T;A)}].
  comp(Glue [psi ⊢→ (T, f)] A)  cT ∈ G ⊢ Compositon(T) supposing G ⊢ (1(𝔽 psi)

Lemma: csm-glue-comp
[G:j⊢]. ∀[K:Top]. ∀[A:{G ⊢ _}]. ∀[psi:{G ⊢ _:𝔽}]. ∀[T:{G, psi ⊢ _}]. ∀[cA,cT,f,tau:Top].
  ((comp(Glue [psi ⊢→ (T, f)] A) )tau comp(Glue [(psi)tau ⊢→ ((T)tau, (f)tau)] (A)tau) )

Lemma: csm-glue-comp-agrees
The type (not displayed) of the equality in this lemma 
is composition-structure{i:l}(G, psi; T)
This means that in "extent" psi, when ⌜G ⊢ Glue [psi ⊢→ (T;f)] T ∈ {G, psi ⊢ _}⌝the
composition for ⌜Glue [psi ⊢→ (T;f)] A⌝ is the same as the composition for T.

This property of the compostion for Glue is used in construction of the 
composition for c𝕌  (the cubiucal universe type).⋅

[H,G:j⊢]. ∀[tau:H j⟶ G]. ∀[A:{G ⊢ _}]. ∀[cA:G +⊢ Compositon(A)]. ∀[psi:{G ⊢ _:𝔽}]. ∀[T:{G, psi ⊢ _}].
[cT:G, psi +⊢ Compositon(T)]. ∀[f:{G, psi ⊢ _:Equiv(T;A)}].
  (comp(Glue [psi ⊢→ (T, f)] A) )tau (cT)tau ∈ H ⊢ Compositon((T)tau) supposing H ⊢ (1(𝔽 (psi)tau)

Definition: gluetype
gluetype(Gamma;A;phi;T;f) ==  Glue [phi ⊢→ (T;equiv-fun(f))] A

Lemma: gluetype_wf
[Gamma:j⊢]. ∀[A:{Gamma ⊢ _}]. ∀[phi:{Gamma ⊢ _:𝔽}]. ∀[T:{Gamma, phi ⊢ _}]. ∀[f:{Gamma, phi ⊢ _:Equiv(T;A)}].
  Gamma ⊢ gluetype(Gamma;A;phi;T;f)

Lemma: glue-comp_wf3
G:j⊢. ∀A:{G ⊢ _}. ∀cA:G +⊢ Compositon(A). ∀psi:{G ⊢ _:𝔽}. ∀T:{G, psi ⊢ _}. ∀cT:G, psi +⊢ Compositon(T).
f:{G, psi ⊢ _:Equiv(T;A)}.
  (comp(Glue [psi ⊢→ (T, f)] A)  ∈ G ⊢ Compositon(gluetype(G;A;psi;T;f)))

Comment: cubical universe doc
============================= cubical universe  ================================

We can now define the cubical universe c𝕌 of fibrant types.
To get cumulativity cubical-universe-cumulativity we must use
the definition of fibrant that quantifies only over representable cubical sets.

compU_wf shows that the Universe has composition structure.⋅

Definition: closed-cubical-universe
cc𝕌 ==  <λI.FibrantType(formal-cube(I)), λI,J,f,FT. csm-fibrant-type(formal-cube(I);formal-cube(J);<f>;FT)>

Lemma: closed-cubical-universe_wf
cc𝕌 ∈ * ⊢_}

Lemma: closed-cubical-universe-cumulativity
closed-cubical-term(cc𝕌) ⊆closed-cubical-term(cc𝕌')

Definition: cubical-universe
c𝕌 ==  closed-type-to-type(cc𝕌)

Lemma: cubical-universe_wf
[X:j⊢]. X ⊢c𝕌

Lemma: cubical-universe-ap-morph
[I,b,J,f,a:Top].  ((b f) csm-fibrant-type(formal-cube(I);formal-cube(J);<f>;b))

Lemma: cubical-term_wf-universe
X:j⊢({X ⊢ _:c𝕌} ∈ 𝕌{[i' j']})

Lemma: istype-cubical-universe-term
X:j⊢istype({X ⊢ _:c𝕌})

Lemma: cubical-universe-at
[I,a:Top].  (c𝕌(a) A:{formal-cube(I) ⊢ _} × formal-cube(I) ⊢ CompOp(A))

Lemma: cube-context-adjoin_wf_universe
[Gamma:⊢]. Gamma.c𝕌 ⊢

Lemma: cubical-universe-at-equal
[X:j⊢]. ∀[I:fset(ℕ)]. ∀[a:X(I)]. ∀[x,y:c𝕌(a)].
  y ∈ c𝕌(a) 
  supposing ((fst(x)) (fst(y)) ∈ {formal-cube(I) ⊢ _}) ∧ ((snd(x)) (snd(y)) ∈ formal-cube(I) ⊢ CompOp(fst(x)))

Definition: universe-type
universe-type(t;I;a) ==  fst(t(a))

Lemma: universe-type_wf
[X:j⊢]. ∀[t:{X ⊢ _:c𝕌}]. ∀[I:fset(ℕ)]. ∀[a:X(I)].  formal-cube(I) ⊢ universe-type(t;I;a)

Lemma: csm-cubical-universe
[s:Top]. ((c𝕌)s c𝕌)

Lemma: csm-ap-term-universe
[X,H:j⊢]. ∀[s:H j⟶ X]. ∀[t:{X ⊢ _:c𝕌}].  ((t)s ∈ {H ⊢ _:c𝕌})

Lemma: cubical-universe-p
[G:j⊢]. ∀[A:{G ⊢ _:c𝕌}]. ∀[T:{G ⊢ _}].  ((A)p ∈ {G.T ⊢ _:c𝕌})

Lemma: cubical-universe-at-cumulativity
[I:fset(ℕ)]. ∀[a:Top].  (c𝕌(a) ⊆c𝕌'(a))

Lemma: cubical-universe-cumulativity2
[X:j⊢]. ∀[t:{X ⊢ _:c𝕌}].  (t ∈ {X ⊢ _:c𝕌'})

Lemma: cubical-universe-cumulativity
[X:j⊢]. ({X ⊢ _:c𝕌} ⊆{X ⊢ _:c𝕌'})

Definition: universe-encode
encode(T;cT) ==  λI,a. <(T)<a>(cT)<a>>

Lemma: universe-encode_wf
[G:j⊢]. ∀[T:{G ⊢ _}]. ∀[cT:G ⊢ CompOp(T)].  (encode(T;cT) ∈ {G ⊢ _:c𝕌})

Lemma: csm-universe-encode
[G:j⊢]. ∀[T:{G ⊢ _}]. ∀[cT:G ⊢ CompOp(T)]. ∀[H:j⊢]. ∀[s:H j⟶ G].  ((encode(T;cT))s encode((T)s;(cT)s) ∈ {H ⊢ _:c𝕌})

Definition: universe-decode
decode(t) ==  <λI,a. fst(t(a))(1), λI,J,f,a,x. (x f)>

Lemma: universe-decode_wf
[X:j⊢]. ∀[t:{X ⊢ _:c𝕌}].  X ⊢ decode(t)

Lemma: csm-universe-decode
[t,s:Top].  ((decode(t))s decode((t)s))

Lemma: universe-decode-type
[X:j⊢]. ∀[t:{X ⊢ _:c𝕌}]. ∀[I:fset(ℕ)]. ∀[rho:X(I)].  (decode((t)<rho>universe-type(t;I;rho) ∈ {formal-cube(I) ⊢ _})

Lemma: universe-decode-restriction
[X:j⊢]. ∀[t:{X ⊢ _:c𝕌}]. ∀[I,J:fset(ℕ)]. ∀[f:I ⟶ J]. ∀[rho:X(J)].
  (decode(t)(f(rho)) universe-type(t;J;rho)(f(1)) ∈ Type)

Lemma: csm-universe-type
[X:j⊢]. ∀[t:{X ⊢ _:c𝕌}]. ∀[I:fset(ℕ)]. ∀[a:X(I)]. ∀[K:fset(ℕ)]. ∀[f:K ⟶ I].
  ((universe-type(t;I;a))<f> universe-type(t;K;f(a)) ∈ {formal-cube(K) ⊢ _})

Lemma: universe-type-at
[X:j⊢]. ∀[t:{X ⊢ _:c𝕌}]. ∀[I:fset(ℕ)]. ∀[a:X(I)]. ∀[K:fset(ℕ)]. ∀[f:K ⟶ I].
  (universe-type(t;I;a)(f) decode(t)(f(a)) ∈ Type)

Definition: universe-comp-op
compOp(t) ==  λI,i,rho. ((snd(t(rho))) 1)

Lemma: universe-comp-op_wf
[X:j⊢]. ∀[t:{X ⊢ _:c𝕌}].  (compOp(t) ∈ X ⊢ CompOp(decode(t)))

Lemma: csm-universe-comp-op
[E,s:Top].  ((compOp(E))s compOp((E)s))

Lemma: universe-encode-decode
[X:j⊢]. ∀[t:{X ⊢ _:c𝕌}].  (encode(decode(t);compOp(t)) t ∈ {X ⊢ _:c𝕌})

Lemma: universe-decode-encode
[X:j⊢]. ∀[T:{X ⊢ _}]. ∀[cT:X ⊢ CompOp(T)].  (decode(encode(T;cT)) T ∈ {X ⊢ _})

Lemma: universe-comp-op-encode
[X:j⊢]. ∀[T:{X ⊢ _}]. ∀[cT:X ⊢ CompOp(T)].  (compOp(encode(T;cT)) cT ∈ X ⊢ CompOp(T))

Definition: universe-comp-fun
CompFun(A) ==  cop-to-cfun(compOp(A))

Lemma: universe-comp-fun_wf
[X:j⊢]. ∀[t:{X ⊢ _:c𝕌}].  (CompFun(t) ∈ +⊢ Compositon(decode(t)))

Lemma: csm-universe-comp-fun
[s,A,G,H:Top].  ((CompFun(A))s CompFun((A)s))

Definition: equiv_comp
equiv_comp(H;A;E;cA;cE) ==
  sigma_comp(pi_comp(H;A;cA;(cE)p);pi_comp(H.(A ⟶ E);(E)p;(cE)p;contractible_comp(H.(A ⟶ E).(E)p;Fiber((q)p;q);
                                                                                   fiber-comp(H.(A ⟶ E).(E)p;((A)p)p;
                                                                                              ((E)p)p;(q)p;q;((cA)p)p;
                                                                                              ((cE)p)p))))

Lemma: equiv_comp_wf
[H:j⊢]. ∀[A,E:{H ⊢ _}]. ∀[cA:H +⊢ Compositon(A)]. ∀[cE:H +⊢ Compositon(E)].
  (equiv_comp(H;A;E;cA;cE) ∈ +⊢ Compositon(Equiv(A;E)))

Lemma: equiv_comp-sq
[H,A,E,cA,cE:Top].
  (equiv_comp(H;A;E;cA;cE) 
  sigma_comp(pi_comp(H;A;cA;λH@0,sigma,phi,u,a0. (cE H@0 x,x@0. (fst((sigma x@0)))) phi a0));
               pi_comp(H.(A ⟶ E);(E)λI,p. (fst(p));λH@0,sigma,phi,u,a0. (cE H@0 x,x@0. (fst((sigma x@0)))) phi a0\000C);
                       contractible_comp(H.(A ⟶ E).(E)λI,p. (fst(p));Fiber(λI,a. (snd(fst(a)));λI,p. (snd(p)));
                                         λH@0,sigma,phi,u,a0,I,a@0.
                                                                   <cA H@0.𝕀 x,x@0. (fst(fst((sigma (m x@0)))))) 
                                                                    I,rho. phi (fst(rho)) ∨ dM-to-FL(I;¬(snd(rho))))\000C 
                                                                    I,rho. if (phi (fst((m rho)))==1)
                                                                            then fst((u (m rho)))
                                                                            else fst((a0 (fst((m rho)))))
                                                                            fi 
                                                                    I,a. (fst((a0 (fst(a)))))) 
                                                                    
                                                                    <a@0, 1>
                                                                   , λJ,f,u@0.
                                                                              (cE H@0.𝕀 
                                                                               x,x@0. (fst(fst((sigma 
                                                                                                  <fst(fst(x@0))
                                                                                                  snd(x@0)
                                                                                                  >))))) 
                                                                               I,rho.
                                                                                       phi (fst(rho)) ∨ ... ∨ (...=1) 
                                                                                                                
                                                                                                                rho) 
                                                                               I,rho.
                                                                                       if (phi (fst(fst(rho)))==1)
                                                                                         then (snd((u 
                                                                                                    <fst(fst(rho))
                                                                                                    snd(rho)
                                                                                                    >))) 
                                                                                              
                                                                                              
                                                                                              (snd(fst(rho)))
                                                                                       if (dM-to-FL(I;¬(snd(...)))==1)
                                                                                         then snd((sigma 
                                                                                                   <fst(fst(rho))
                                                                                                   snd(rho)
                                                                                                   >))
                                                                                       else (snd(fst((sigma 
                                                                                                      <fst(fst(rho))
                                                                                                      snd(rho)
                                                                                                      >)))) 
                                                                                            
                                                                                            
                                                                                            (cA H@0.𝕀 
                                                                                             x,x@0.
                                                                                                     (fst(fst((sigma 
                                                                                                               (m 
                                                                                                                x@0)))))\000C) 
                                                                                             I,a.
                                                                                                   (phi 
                                                                                                    (fst(a))) ∨ I,p. .\000C..=0)) 
                                                                                             I,rho.
                                                                                                     if (...==1)
                                                                                                     then fst((u 
                                                                                                               (m 
                                                                                                                rho)))
                                                                                                     else fst((a0 
                                                                                                               ...))
                                                                                                     fi 
                                                                                             I,a. (fst((a0 
                                                                                                          (fst(a)))))) 
                                                                                             
                                                                                             <fst(fst(rho)), snd(rho)>)
                                                                                       fi 
                                                                               I,a. ((snd((a0 (fst(a))))) 
                                                                                       (snd(a)))) 
                                                                               
                                                                               (f(a@0);u@0))
                                                                   >))))

Definition: equiv_comp_apply
equiv_comp_apply(H;A;E;cA;cE;I;a;b;c;d) ==
  <λJ,f,u@0. (cE <λJ.J ⟶ I, λI,J,f,g. g ⋅ f>.𝕀.((A)λx,x@0. <a[x;m x@0], b[x;m x@0]>)[1(𝕀)] 
              x,x@0. <a[x;m x <fst(fst(x@0)), snd(x@0)>], b[x;m x <fst(fst(x@0)), snd(x@0)>]>
              I,a. c[I;fst(fst(a))] ∨ dM-to-FL(I;¬(snd(fst(a))))) 
              I@0,a@0. (if (c[I@0;fst((m I@0 <fst(fst(a@0)), snd(a@0)>))]==1)
                          then fst(⋅)
                          else fst(d[I@0;fst((m I@0 <fst(fst(a@0)), snd(a@0)>))])
                          fi  
                          I@0 
                          
                          (cA <λJ.J ⟶ I, λI,J,f,g. g ⋅ f>.𝕀.((A)λx,x@0. <a[x;m x@0], b[x;m x@0]>)[1(𝕀)].𝕀 
                           x,x@0. <a[x;m x <fst(fst((m x@0))), ¬(snd((m x@0)))>]
                                    b[x;m x <fst(fst((m x@0))), ¬(snd((m x@0)))>]
                                    >
                           I,rho. 0 ∨ dM-to-FL(I;¬(snd(rho)))) 
                           I,rho. (snd(fst((m rho))))) 
                           I,a. (snd(fst(a)))) 
                           I@0 
                           <fst(a@0), ¬(snd(a@0))>))) 
              I@0,a@0. ((fst(d[I@0;fst(fst(a@0))])) I@0 
                          (cA <λJ.J ⟶ I, λI,J,f,g. g ⋅ f>.𝕀.((A)λx,x@0. <a[x;m x@0], b[x;m x@0]>)[1(𝕀)].𝕀 
                           x,x@0. <a[x;m x <fst(fst((m x@0))), ¬(snd((m x@0)))>]
                                    b[x;m x <fst(fst((m x@0))), ¬(snd((m x@0)))>]
                                    >
                           I,rho. 0 ∨ dM-to-FL(I;¬(snd(rho)))) 
                           I,rho. (snd(fst((m rho))))) 
                           I,a. (snd(fst(a)))) 
                           I@0 
                           <a@0, ¬(0)>))) 
              
              (f(<1, 1>);u@0))
  , λJ,f,u@0.
             (cubical-pair(λI@0,a@1.
                                    <cA 
                                     <λJ.J ⟶ I, λI,J,f,g. g ⋅ f>.(let x,F 
                                                           in <λJ,c. (x J <a[J;c], b[J;c]>)
                                                              , λK,J,f,c,u. (F f <a[K;c], b[K;c]> u)
                                                              >)[1(𝕀)].𝕀.𝕀 
                                     x,x@0. <a[x;<fst(fst((m (m x@0)))), snd((m (m x@0)))>]
                                              b[x;<fst(fst((m (m x@0)))), snd((m (m x@0)))>]
                                              >
                                     I,rho.
                                             c[I;
                                              fst(fst(fst(rho)))] ∨ dM-to-FL(I;¬(snd(fst(rho)))) ∨ dM-to-FL(I;¬(...))) 
                                     I@0,rho.
                                               if (c[I@0;fst(fst(fst((m I@0 
                                                                      rho))))] ∨ dM-to-FL(I@0;¬(snd(fst((m I@0 
                                                                                                         rho)))))==1)
                                               then fst(if (c[I@0;fst(fst((m I@0 (m I@0 rho))))]==1)
                                                    then fst(((snd(⋅)) I@0 
                                                              (cE 
                                                               <λJ.J ⟶ I, λI,J,f,g. g ⋅ f>.(let x,F 
                                                                                     in <λJ,c. (x J <a[J;c], b[J;c]>)
                                                                                        , λK,J,f,c,u. (F 
                                                                                                       <a[K;c], b[K;c]> 
                                                                                                       u)
                                                                                        >)[1(𝕀)].𝕀 
                                                               x,x@0. <a[x;<fst(fst((m x@0))), ¬(snd((m x@0)))>]
                                                                        b[x;<fst(fst((m x@0))), ¬(snd((m x@0)))>]
                                                                        >
                                                               I,rho. 0 ∨ dM-to-FL(I;¬(snd(rho)))) 
                                                               I,rho. (snd(fst((m rho))))) 
                                                               I,a. (snd(fst(a)))) 
                                                               I@0 
                                                               <fst((m I@0 (m I@0 rho)))
                                                               , ¬(snd((m I@0 (m I@0 rho))))
                                                               >)))
                                                    else fst(...)
                                                    fi )
                                               else ...
                                               fi 
                                     ... 
                                     ... 
                                     ...
                                    ...
                                    >;...) 
              ... 
              ...)
  >

Lemma: equiv_comp-apply-sq
[H,A,E,cA,cE,I,a,b,c,d:Top].
  (equiv_comp(H;A;E;cA;cE) formal-cube(I) x,y. <a[x;y], b[x;y]>K,f. c[K;f]) I@0,a. ⋅K,f. d[K;f]) 
  equiv_comp_apply(H;A;E;cA;cE;I;a;b;c;d))

Lemma: csm-equiv_comp-sq
[H,K,tau,A,E,cA,cE:Top].
  ((equiv_comp(H;A;E;cA;cE))tau sigma_comp(pi_comp(K;(A)tau;(cA)tau;((cE)tau)p);
                                             pi_comp(K.((A)tau ⟶ (E)tau);((E)tau)p;((cE)tau)p;
                                                     contractible_comp(K.((A)tau ⟶ (E)tau).((E)tau)p;
                                                                       (Fiber((q)p;q))tau++;
                                                                       fiber-comp(K.((A)tau ⟶ (E)tau).((E)tau)p;
                                                                                  (((A)tau)p)p;(((E)tau)p)p;(q)p;q;
                                                                                  (((cA)tau)p)p;(((cE)tau)p)p)))))

Lemma: csm-equiv_comp
[H,K:j⊢]. ∀[tau:K j⟶ H]. ∀[A,E:{H ⊢ _}]. ∀[cA:H +⊢ Compositon(A)]. ∀[cE:H +⊢ Compositon(E)].
  ((equiv_comp(H;A;E;cA;cE))tau equiv_comp(K;(A)tau;(E)tau;(cA)tau;(cE)tau) ∈ K ⊢ Compositon(Equiv((A)tau;(E)tau)))

Definition: equiv-comp
equiv-comp(H;A;E;cA;cE) ==  cfun-to-cop(H;Equiv(A;E);equiv_comp(H;A;E;cop-to-cfun(cA);cop-to-cfun(cE)))

Lemma: equiv-comp_wf
[H:j⊢]. ∀[A,E:{H ⊢ _}]. ∀[cA:H ⊢ CompOp(A)]. ∀[cE:H ⊢ CompOp(E)].  (equiv-comp(H;A;E;cA;cE) ∈ H ⊢ CompOp(Equiv(A;E)))

Lemma: equiv-comp-exists
H:j⊢. ∀A,E:{H ⊢ _}.  (H ⊢ CompOp(A)  H ⊢ CompOp(E)  H ⊢ CompOp(Equiv(A;E)))

Lemma: csm-equiv-comp
[H,K:j⊢]. ∀[tau:K j⟶ H]. ∀[A,E:{H ⊢ _}]. ∀[cA:H ⊢ CompOp(A)]. ∀[cE:H ⊢ CompOp(E)].
  ((equiv-comp(H;A;E;cA;cE))tau equiv-comp(K;(A)tau;(E)tau;(cA)tau;(cE)tau) ∈ K ⊢ CompOp(Equiv((A)tau;(E)tau)))

Definition: equivU
equivU(G;E;cE) ==  transport(G;IdEquiv(G;(E)[0(𝕀)]))

Lemma: equivU_wf
[G:j⊢]. ∀[E:{G.𝕀 ⊢ _}]. ∀[cE:G.𝕀 ⊢ CompOp(E)].  (equivU(G;E;cE) ∈ {G ⊢ _:Equiv((E)[0(𝕀)];(E)[1(𝕀)])})

Definition: comp-U
comp-U() ==
  λG,phi,E,A. encode(Glue [phi ⊢→ (decode((E)[1(𝕀)]);equiv-fun(equivU(G, phi;(decode(E))-;(compOp(E))-)))] decode(A);
                     cfun-to-cop(G;Glue [phi ⊢→ (decode((E)[1(𝕀)]);equiv-fun(equivU(G, phi;(decode(E))-;(compOp(E))-)))]
                                        decode(A)
                         ;comp(Glue [phi ⊢→ (decode((E)[1(𝕀)]), equivU(G, phi;(decode(E))-;(compOp(E))-))] decode(A)) ))

Lemma: csm-equivU
[G,K:j⊢]. ∀[tau:K j⟶ G]. ∀[E:{G.𝕀 ⊢ _}]. ∀[cE:G.𝕀 ⊢ CompOp(E)].
  ((equivU(G;E;cE))tau equivU(K;(E)tau+;(cE)tau+) ∈ {K ⊢ _:Equiv(((E)tau+)[0(𝕀)];((E)tau+)[1(𝕀)])})

Definition: compU
compU() ==
  λG,s,phi,E,A. encode(Glue [phi ⊢→ (decode((E)[1(𝕀)]);equiv-fun(equivU(G, phi;(decode(E))-;(compOp(E))-)))] decode(A);
                       cfun-to-cop(G;Glue [phi ⊢→ (decode((E)[1(𝕀)]);equiv-fun(equivU(G, phi;(decode(E))-;
                                                                                      (compOp(E))-)))] decode(A)
                           ;comp(Glue [phi ⊢→ (decode((E)[1(𝕀)]), equivU(G, phi;(decode(E))-;(compOp(E))-))]
                                     decode(A)) ))

Lemma: compU-wf-lemma1
[G:j⊢]. ∀[phi:{G ⊢ _:𝔽}].
  E,A. encode(Glue [phi ⊢→ (decode((E)[1(𝕀)]);equiv-fun(equivU(G, phi;(decode(E))-;(compOp(E))-)))] decode(A);
                cfun-to-cop(G;Glue [phi ⊢→ (decode((E)[1(𝕀)]);equiv-fun(equivU(G, phi;(decode(E))-;(compOp(E))-)))]
                                   decode(A)
                    ;comp(Glue [phi ⊢→ (decode((E)[1(𝕀)]), equivU(G, phi;(decode(E))-;(compOp(E))-))] decode(A)) ))
   ∈ u:{G, phi.𝕀 ⊢ _:c𝕌} ⟶ a0:{G ⊢ _:c𝕌[phi |⟶ (u)[0(𝕀)]]} ⟶ {G ⊢ _:c𝕌[phi |⟶ (u)[1(𝕀)]]})

Lemma: compU_wf1
[G:j⊢]. (compU() ∈ composition-function{j:l, i':l}(G; c𝕌))

Lemma: compU_wf
[G:j⊢]. (compU() ∈ G ⊢ Compositon'(c𝕌))

Definition: univ-comp
univ-comp{i:l}() ==  cfun-to-cop(();c𝕌;compU())

Lemma: univ-comp-sq
[G:Top]. (univ-comp{i:l}() cfun-to-cop(G;c𝕌;compU()))

Lemma: univ-comp_wf
[G:j⊢]. (univ-comp{i:l}() ∈ G ⊢ CompOp(c𝕌))

Lemma: comp-op-to-comp-fun-univ-comp
[G:j⊢]. (cop-to-cfun(univ-comp{i:l}()) ∈ composition-structure{[i' j]:l, i':l}(G; c𝕌))

Lemma: csm-univ-comp
[s:Top]. ((univ-comp{i:l}())s univ-comp{i:l}())

Lemma: composition-op-universe-sq
[G:Top]
  (G ⊢ CompOp(c𝕌{comp:I:fset(ℕ)
                     ⟶ i:{i:ℕ| ¬i ∈ I} 
                     ⟶ rho:G(I+i)
                     ⟶ phi:𝔽(I)
                     ⟶ u:{I+i,s(phi) ⊢ _:c𝕌}
                     ⟶ {a0:FibrantType(formal-cube(I))| 
                         ∀J:fset(ℕ). ∀f:I,phi(J).
                           (csm-fibrant-type(formal-cube(I);formal-cube(J);<f>;a0)
                           u((i0) ⋅ f)
                           ∈ FibrantType(formal-cube(J)))} 
                     ⟶ {a1:FibrantType(formal-cube(I))| 
                         ∀J:fset(ℕ). ∀f:I,phi(J).
                           (csm-fibrant-type(formal-cube(I);formal-cube(J);<f>;a1)
                           u((i1) ⋅ f)
                           ∈ FibrantType(formal-cube(J)))} 
                     ∀I,J:fset(ℕ). ∀i:{i:ℕ| ¬i ∈ I} . ∀j:{j:ℕ| ¬j ∈ J} . ∀g:J ⟶ I. ∀rho:G(I+i). ∀phi:𝔽(I).
                     ∀u:{I+i,s(phi) ⊢ _:c𝕌}. ∀a0:{a0:FibrantType(formal-cube(I))| 
                                                  ∀J:fset(ℕ). ∀f:I,phi(J).
                                                    (csm-fibrant-type(formal-cube(I);formal-cube(J);<f>;a0)
                                                    u((i0) ⋅ f)
                                                    ∈ FibrantType(formal-cube(J)))} .
                       (csm-fibrant-type(formal-cube(I);formal-cube(J);<g>;comp rho phi a0)
                       (comp g,i=j(rho) g(phi) (u)subset-trans(I+i;J+j;g,i=j;s(phi)) 
                          csm-fibrant-type(formal-cube(I);formal-cube(J);<g>;a0))
                       ∈ FibrantType(formal-cube(J)))} )

Definition: universe-term
t𝕌 ==  encode(c𝕌;univ-comp{i:l}())

Lemma: universe-term_wf
[G:j⊢]. (t𝕌 ∈ {G ⊢ _:c𝕌'})

Lemma: decode-universe-term
[G:j⊢]. (decode(t𝕌c𝕌 ∈ {G ⊢_})

Lemma: comp-universe-term
[G:j⊢]. (compOp(t𝕌univ-comp{i:l}() ∈ G ⊢ CompOp(c𝕌))

Definition: cubical-equiv-by-cases
cubical-equiv-by-cases(G;B;f) ==  ((f)p ∨ IdEquiv(G.𝕀(q=1);(B)p))

Lemma: cubical-equiv-by-cases_wf
[G:j⊢]. ∀[A,B:{G ⊢ _}]. ∀[f:{G ⊢ _:Equiv(A;B)}].
  (cubical-equiv-by-cases(G;B;f) ∈ {G.𝕀((q=0) ∨ (q=1)) ⊢ _:Equiv((if (q=0) then (A)p else (B)p);(B)p)})

Lemma: csm-cubical-equiv-by-cases
G:j⊢. ∀A,B:{G ⊢ _}. ∀f:{G ⊢ _:Equiv(A;B)}. ∀H:j⊢. ∀s:H j⟶ G.
  ((cubical-equiv-by-cases(G;B;f))s+
  cubical-equiv-by-cases(H;(B)s;(f)s)
  ∈ {H.𝕀((q=0) ∨ (q=1)) ⊢ _:Equiv((if (q=0) then ((A)s)p else ((B)s)p);((B)s)p)})

Definition: equiv-path1
equiv-path1(G;A;B;f) ==
  Glue [((q=0) ∨ (q=1)) ⊢→ ((if (q=0) then (A)p else (B)p);equiv-fun(cubical-equiv-by-cases(G;B;f)))] (B)p

Lemma: equiv-path1_wf
[G:j⊢]. ∀[A,B:{G ⊢ _}]. ∀[f:{G ⊢ _:Equiv(A;B)}].  G.𝕀 ⊢ equiv-path1(G;A;B;f)

Lemma: csm-equiv-path1
[G:j⊢]. ∀[A,B:{G ⊢ _}]. ∀[f:{G ⊢ _:Equiv(A;B)}]. ∀[H:j⊢]. ∀[s:H j⟶ G].
  ((equiv-path1(G;A;B;f))s+ equiv-path1(H;(A)s;(B)s;(f)s) ∈ {H.𝕀 ⊢ _})

Lemma: equiv-path1-constraint
[G:j⊢]. ∀[A,B:{G ⊢ _}]. ∀[f:{G ⊢ _:Equiv(A;B)}].
  G.𝕀((q=0) ∨ (q=1)) ⊢ equiv-path1(G;A;B;f) (if (q=0) then (A)p else (B)p)

Lemma: equiv-path1-0
[G:j⊢]. ∀[A,B:{G ⊢ _}]. ∀[f:{G ⊢ _:Equiv(A;B)}].  ((equiv-path1(G;A;B;f))[0(𝕀)] A ∈ {G ⊢ _})

Lemma: equiv-path1-1
[G:j⊢]. ∀[A,B:{G ⊢ _}]. ∀[f:{G ⊢ _:Equiv(A;B)}].  ((equiv-path1(G;A;B;f))[1(𝕀)] B ∈ {G ⊢ _})

Lemma: face-term-0-and-1
G:j⊢G.𝕀 ⊢ (((q=0) ∧ (q=1))  0(𝔽))

Definition: equiv-path2
equiv-path2(G;A;B;cA;cB;f) ==
  comp(Glue [((q=0) ∨ (q=1)) ⊢→ ((if (q=0) then (A)p else (B)p), cubical-equiv-by-cases(G;B;f))] (B)p) 

Lemma: equiv-path2_wf
[G:j⊢]. ∀[A,B:{G ⊢ _}]. ∀[cA:G +⊢ Compositon(A)]. ∀[cB:G +⊢ Compositon(B)]. ∀[f:{G ⊢ _:Equiv(A;B)}].
  (equiv-path2(G;A;B;cA;cB;f) ∈ G.𝕀 +⊢ Compositon(equiv-path1(G;A;B;f)))

Lemma: csm-equiv-path2
[G:j⊢]. ∀[A,B:{G ⊢ _}]. ∀[cA:G +⊢ Compositon(A)]. ∀[cB:G +⊢ Compositon(B)]. ∀[f:{G ⊢ _:Equiv(A;B)}]. ∀[H:j⊢].
[s:H j⟶ G].
  ((equiv-path2(G;A;B;cA;cB;f))s+
  equiv-path2(H;(A)s;(B)s;(cA)s;(cB)s;(f)s)
  ∈ H.𝕀 +⊢ Compositon(equiv-path1(H;(A)s;(B)s;(f)s)))

Lemma: equiv-path2-0
[G:j⊢]. ∀[A,B:{G ⊢ _}]. ∀[f:{G ⊢ _:Equiv(A;B)}]. ∀[cA:G +⊢ Compositon(A)]. ∀[cB:G +⊢ Compositon(B)].
  ((equiv-path2(G;A;B;cA;cB;f))[0(𝕀)] cA ∈ +⊢ Compositon(A))

Lemma: equiv-path2-1
[G:j⊢]. ∀[A,B:{G ⊢ _}]. ∀[f:{G ⊢ _:Equiv(A;B)}]. ∀[cA:G +⊢ Compositon(A)]. ∀[cB:G +⊢ Compositon(B)].
  ((equiv-path2(G;A;B;cA;cB;f))[1(𝕀)] cB ∈ +⊢ Compositon(B))

Definition: equiv_path
equiv_path(G;A;B;f) ==
  let equiv-path1(G;decode(A);decode(B);f) in
   let equiv-path2(G;decode(A);decode(B);CompFun(A);CompFun(B);f) in
   encode(T;cfun-to-cop(G.𝕀;T;C))

Lemma: equiv_path_wf
[G:j⊢]. ∀[A,B:{G ⊢ _:c𝕌}]. ∀[f:{G ⊢ _:Equiv(decode(A);decode(B))}].  (equiv_path(G;A;B;f) ∈ {G.𝕀 ⊢ _:c𝕌})

Lemma: csm-equiv_path
[G:j⊢]. ∀[A,B:{G ⊢ _:c𝕌}]. ∀[f:{G ⊢ _:Equiv(decode(A);decode(B))}]. ∀[H:j⊢]. ∀[s:H j⟶ G].
  ((equiv_path(G;A;B;f))s+ equiv_path(H;(A)s;(B)s;(f)s) ∈ {H.𝕀 ⊢ _:c𝕌})

Lemma: equiv_path-0
[G:j⊢]. ∀[A,B:{G ⊢ _:c𝕌}]. ∀[f:{G ⊢ _:Equiv(decode(A);decode(B))}].  G ⊢ (equiv_path(G;A;B;f))[0(𝕀)]=A:c𝕌

Lemma: equiv_path-1
[G:j⊢]. ∀[A,B:{G ⊢ _:c𝕌}]. ∀[f:{G ⊢ _:Equiv(decode(A);decode(B))}].  G ⊢ (equiv_path(G;A;B;f))[1(𝕀)]=B:c𝕌

Definition: equiv-path
EquivPath(G;A;B;f) ==  <>(equiv_path(G;A;B;f))

Lemma: equiv-path_wf
[G:j⊢]. ∀[A,B:{G ⊢ _:c𝕌}]. ∀[f:{G ⊢ _:Equiv(decode(A);decode(B))}].  (EquivPath(G;A;B;f) ∈ {G ⊢ _:(Path_c𝕌 B)})

Definition: univ-a
UA ==  cubical-lam(G;EquivPath(G.Equiv(decode(A);decode(B));(A)p;(B)p;q))

Lemma: univ-a_wf
[G:j⊢]. ∀[A,B:{G ⊢ _:c𝕌}].  (UA ∈ {G ⊢ _:(Equiv(decode(A);decode(B)) ⟶ (Path_c𝕌 B))})

Lemma: univ-a_wf2
[G:j⊢]. (UA ∈ {G.c𝕌.c𝕌 ⊢ _:(Equiv(decode(q);decode((q)p)) ⟶ (Path_c𝕌 (q)p))})

Lemma: app-univ-a-1
[G:j⊢]. ∀[A,B:{G ⊢ _:c𝕌}]. ∀[f:{G ⊢ _:Equiv(decode(A);decode(B))}].
  (app(UA; f) (EquivPath(G.Equiv(decode(A);decode(B));(A)p;(B)p;q))[f] ∈ {G ⊢ _:(Path_c𝕌 B)})

Lemma: app-univ-a
[G:j⊢]. ∀[A,B:{G ⊢ _:c𝕌}]. ∀[f:{G ⊢ _:Equiv(decode(A);decode(B))}].
  (app(UA; f) EquivPath(G;A;B;f) ∈ {G ⊢ _:(Path_c𝕌 B)})

Definition: univ-trans
univ-trans(G;T) ==  transprt-fun(G;decode(T);cop-to-cfun(compOp(T)))

Lemma: univ-trans_wf
[G:j⊢]. ∀[T:{G.𝕀 ⊢ _:c𝕌}].  (univ-trans(G;T) ∈ {G ⊢ _:((decode(T))[0(𝕀)] ⟶ (decode(T))[1(𝕀)])})

Lemma: csm-univ-trans
[G:j⊢]. ∀[T:{G.𝕀 ⊢ _:c𝕌}]. ∀[H:j⊢]. ∀[s:H j⟶ G].
  ((univ-trans(G;T))s univ-trans(H;(T)s+) ∈ {H ⊢ _:(((decode(T))s+)[0(𝕀)] ⟶ ((decode(T))s+)[1(𝕀)])})

Definition: path-trans
PathTransport(p) ==  univ-trans(G;path-eta(p))

Lemma: path-trans_wf
[G:j⊢]. ∀[A,B:{G ⊢ _:c𝕌}]. ∀[p:{G ⊢ _:(Path_c𝕌 B)}].  (PathTransport(p) ∈ {G ⊢ _:(decode(A) ⟶ decode(B))})

Lemma: path-trans-sq
[G,pth,I,a,J,h,u:Top].  (PathTransport(pth) compOp(path-eta(pth)) new-name(J) <s(h(a)), <new-name(J)>> \000Cdiscr(⋅u)

Lemma: path-trans-sq2
[G,pth,I,a,J,h,u:Top].
  (PathTransport(pth) (snd((pth(s(h(a))) J+new-name(J) 1 <new-name(J)>))) new-name(J) discr(⋅u)

Definition: cubical-subst
cubical-subst(G;f;pth;x) ==  app(PathTransport(map-path(G;f;pth)); x)

Lemma: cubical-subst_wf
[G:j⊢]. ∀[A:{G ⊢_}]. ∀[a,b:{G ⊢ _:A}]. ∀[p:{G ⊢ _:(Path_A b)}]. ∀[f:{G ⊢ _:(A ⟶ c𝕌)}].
[x:{G ⊢ _:decode(app(f; a))}].
  (cubical-subst(G;f;p;x) ∈ {G ⊢ _:decode(app(f; b))})

Definition: equivTerm
equivTerm(G;A;B) ==  encode(Equiv(decode(A);decode(B));equiv-comp(G;decode(A);decode(B);compOp(A);compOp(B)))

Lemma: equivTerm_wf
[G:j⊢]. ∀[A,B:{G ⊢ _:c𝕌}].  (equivTerm(G;A;B) ∈ {G ⊢ _:c𝕌})

Lemma: csm-equivTerm
[G:j⊢]. ∀[A,B:{G ⊢ _:c𝕌}]. ∀[H:j⊢]. ∀[s:H j⟶ G].  ((equivTerm(G;A;B))s equivTerm(H;(A)s;(B)s) ∈ {H ⊢ _:c𝕌})

Lemma: decode-equivTerm
[G:j⊢]. ∀[A,B:{G ⊢ _:c𝕌}].  (decode(equivTerm(G;A;B)) Equiv(decode(A);decode(B)) ∈ {G ⊢ _})

Definition: transEquiv
transEquiv{i:l}(G;A;p) ==  cubical-subst(G;cubical-lam(G;equivTerm(G.c𝕌;(A)p;q));p;IdEquiv(G;decode(A)))

Lemma: transEquiv_wf
[G:j⊢]. ∀[A,B:{G ⊢ _:c𝕌}]. ∀[p:{G ⊢ _:(Path_c𝕌 B)}].  (transEquiv{i:l}(G;A;p) ∈ {G ⊢ _:Equiv(decode(A);decode(B))})

Definition: transEquiv-trans
transEquivFun(p) ==  equiv-fun(transEquiv{i:l}(G;A;p))

Lemma: transEquiv-trans_wf
[G:j⊢]. ∀[A,B:{G ⊢ _:c𝕌}]. ∀[p:{G ⊢ _:(Path_c𝕌 B)}].  (transEquivFun(p) ∈ {G ⊢ _:(decode(A) ⟶ decode(B))})

Lemma: transEquiv-trans-sq
[G,A,p,I,a,J,f,u:Top].
  (transEquivFun(p) let x,cA I+new-name(I) 1(s(1(a))) I+new-name(I) 1 <new-name(I)> 
                                in cA new-name(J) 
                                   1 ⋅ cube+(I;new-name(I)) J+new-name(J) <1 ⋅ f ⋅ s, 1 ∧ <new-name(J)>> ⋅ 
                                   (0)<1 ⋅ f> ∨ 
                                   I@1,a@0. (if ((0)<1 ⋅ f ⋅ s ⋅ a@0>==1) then (fst(⋅)) I@1 else λu.u fi  
                                               ((snd((A I@1+new-name(I@1) 
                                                      cube+(I;new-name(I)) I@1+new-name(I@1) 
                                                      <1 ⋅ f ⋅ s ⋅ a@0 ⋅ s
                                                      1 ∧ ¬(dM-lift(I@1+new-name(I@1);I@1;s) 
                                                              ¬(dM-lift(I@1;J+new-name(J);a@0) 
                                                                <new-name(J)>) ∧ <new-name(I@1)>)
                                                      >(1(1(1(s(1(a))))))))) 
                                                I@1 
                                                new-name(I@1) 
                                                
                                                0 ∨ dM-to-FL(I@1;¬(dM-lift(I@1;J+new-name(J);a@0) <new-name(J)>))) 
                                                I@0,a@1. ((((u s) a@0) s) a@1)) 
                                                ((u s) a@0)))) 
                                   ((snd((A J+new-name(J) 
                                          cube+(I;new-name(I)) J+new-name(J) <1 ⋅ f ⋅ s, 1 ∧ ¬(1 ∧ <new-name(J)>)>(1(1(1\000C(s(1(a))))))))) 
                                    
                                    new-name(J) 
                                    
                                    
                                    I@1,a@0. ((u s) a@0)) 
                                    u))

Lemma: universe-path-type-lemma-0
G:j⊢. ∀A,B:{G ⊢ _:c𝕌}. ∀p:{G ⊢ _:(Path_c𝕌 B)}. ∀I,J:fset(ℕ). ∀f:J ⟶ I. ∀v:G(I+new-name(I)).
  (universe-type(A;I+new-name(I);v)((new-name(I)0) ⋅ f)
  fst((p(v) I+new-name(I) 1 <new-name(I)>))((new-name(I)0) ⋅ f)
  ∈ Type)

Lemma: universe-path-type-lemma-1
G:j⊢. ∀A,B:{G ⊢ _:c𝕌}. ∀p:{G ⊢ _:(Path_c𝕌 B)}. ∀I,J:fset(ℕ). ∀f:J ⟶ I. ∀v:G(I+new-name(I)).
  (universe-type(B;I+new-name(I);v)((new-name(I)1) ⋅ f)
  fst((p(v) I+new-name(I) 1 <new-name(I)>))((new-name(I)1) ⋅ f)
  ∈ Type)

Lemma: transEquiv-trans-eq
[G:j⊢]. ∀[A,B:{G ⊢ _:c𝕌}]. ∀[p:{G ⊢ _:(Path_c𝕌 B)}].
  (transEquivFun(p)
  I,a,J,f,u. ((snd((p(s(a)) I+new-name(I) 1 <new-name(I)>))) new-name(J) f,new-name(I)=new-name(J) 0 ⋅ 
                  (compOp(A) new-name(J) s(f(a)) 0 ⋅ u)))
  ∈ {G ⊢ _:(decode(A) ⟶ decode(B))})

Lemma: transEquiv-trans-eq2
[G:j⊢]. ∀[A,B:{G ⊢ _:c𝕌}]. ∀[p:{G ⊢ _:(Path_c𝕌 B)}].
  (transEquivFun(p)
  I,a,J,h,u. ((snd((p(s(h(a))) J+new-name(J) 1 <new-name(J)>))) new-name(J) discr(⋅
                  (compOp(A) new-name(J) s(h(a)) 0 ⋅ u)))
  ∈ {G ⊢ _:(decode(A) ⟶ decode(B))})

Lemma: transEquiv-trans-eq-path-trans
[G:j⊢]. ∀[A,B:{G ⊢ _:c𝕌}]. ∀[p:{G ⊢ _:(Path_c𝕌 B)}].
  (transEquivFun(p) (PathTransport(p) ConstTrans(decode(A))) ∈ {G ⊢ _:(decode(A) ⟶ decode(B))})

Definition: trans-equiv-path
trans-equiv-path(G;A;B;f) ==
  let tr = λb.transprt-const(G.decode(A);(CompFun(B))p;b) in
   let app(equiv-fun((f)p); q) in
   cubical-lam(G;tr (tr b))

Lemma: trans-equiv-path_wf
[G:j⊢]. ∀[A,B:{G ⊢ _:c𝕌}]. ∀[f:{G ⊢ _:Equiv(decode(A);decode(B))}].
  (trans-equiv-path(G;A;B;f) ∈ {G ⊢ _:(decode(A) ⟶ decode(B))})

Lemma: csm-trans-equiv-path
[G:j⊢]. ∀[A,B:{G ⊢ _:c𝕌}]. ∀[f:{G ⊢ _:Equiv(decode(A);decode(B))}]. ∀[H:j⊢]. ∀[s:H j⟶ G].
  ((trans-equiv-path(G;A;B;f))s trans-equiv-path(H;(A)s;(B)s;(f)s) ∈ {H ⊢ _:(decode((A)s) ⟶ decode((B)s))})

Lemma: app-trans-equiv-path
[G:j⊢]. ∀[A,B:{G ⊢ _:c𝕌}]. ∀[f:{G ⊢ _:Equiv(decode(A);decode(B))}]. ∀[a:{G ⊢ _:decode(A)}].
  (app(trans-equiv-path(G;A;B;f); a)
  transprt-const(G;CompFun(B);transprt-const(G;CompFun(B);app(equiv-fun(f); a)))
  ∈ {G ⊢ _:decode(B)})

Lemma: univ-trans-equiv_path
[G:j⊢]. ∀[A,B:{G ⊢ _:c𝕌}]. ∀[f:{G ⊢ _:Equiv(decode(A);decode(B))}].
  (univ-trans(G;equiv_path(G;A;B;f)) trans-equiv-path(G;A;B;f) ∈ {G ⊢ _:(decode(A) ⟶ decode(B))})

Definition: univalence
Univalence ==  Πc𝕌 Contractible(Σ c𝕌 Equiv(decode((q)p);decode(q)))

Lemma: univalence_wf
[G:j⊢]. G ⊢Univalence

Definition: uabeta_aux
uabeta_aux(G;A;B;f) ==
  let app(equiv-fun((f)p); q) in
   let b' transprt-const(G.decode(A);(CompFun(B))p;b) in
   let pth1 trans-const-path(G.decode(A);(CompFun(B))p;b') in
   let pth2 trans-const-path(G.decode(A);(CompFun(B))p;b) in
   pth1 pth2

Lemma: uabeta_aux_wf
[G:j⊢]. ∀[A,B:{G ⊢ _:c𝕌}]. ∀[f:{G ⊢ _:Equiv(decode(A);decode(B))}].
  (uabeta_aux(G;A;B;f) ∈ {G.decode(A) ⊢ _:let app(equiv-fun((f)p); q) in
                                           let b' transprt-const(G.decode(A);(CompFun(B))p;b) in
                                           let b'' transprt-const(G.decode(A);(CompFun(B))p;b') in
                                           (Path_(decode(B))p b'' b)})

Definition: transport-type
TransportType(A) ==  ∀[B:{G ⊢ _:c𝕌}]. ({G ⊢ _:(Path_c𝕌 B)} ⟶ {G ⊢ _:(decode(A) ⟶ decode(B))})

Lemma: transport-type_wf
[G:j⊢]. ∀[A:{G ⊢ _:c𝕌}].  (TransportType(A) ∈ 𝕌{[i'' j'']})

Definition: uabetatype
uabetatype(G;A;B;f) ==
  ΠEquiv(decode(A);decode(B)) Π(decode(A))p (Path_((decode(B))p)p app(f G.Equiv(decode(A);decode(B)).(decode(A))p 
                                                                      ((A)p)p 
                                                                      app(UA; (q)p); q) app(equiv-fun((q)p); q))

Lemma: uabetatype_wf
[G:j⊢]. ∀[A,B:{G ⊢ _:c𝕌}]. ∀[f:G:CubicalSet{i|j} ⟶ A:{G ⊢ _:c𝕌} ⟶ TransportType(A)].  G ⊢ uabetatype(G;A;B;f)

Definition: uabeta-type
uabeta-type(G;A;B) ==
  ΠEquiv(decode(A);decode(B)) Π(decode(A))p (Path_((decode(B))p)p app(PathTransport(app(UA; (q)p)); q)
                                                  app(equiv-fun((q)p); q))

Lemma: uabeta-type_wf
[G:j⊢]. ∀[A,B:{G ⊢ _:c𝕌}].  G ⊢ uabeta-type(G;A;B)

Definition: uabeta
uabeta(G;A;B) ==  uabeta_aux(G.Equiv(decode(A);decode(B));(A)p;(B)p;q)))

Lemma: uabeta_wf
[G:j⊢]. ∀[A,B:{G ⊢ _:c𝕌}].  (uabeta(G;A;B) ∈ {G ⊢ _:uabeta-type(G;A;B)})

Definition: transEquivbeta-type
transEquivbeta-type{i:l}(G;A;B) ==
  ΠEquiv(decode(A);decode(B)) Π(decode(A))p (Path_((decode(B))p)p app(transEquivFun(app(UA; (q)p)); q)
                                                  app(equiv-fun((q)p); q))

Lemma: transEquivbeta-type_wf
[G:j⊢]. ∀[A,B:{G ⊢ _:c𝕌}].  G ⊢ transEquivbeta-type{i:l}(G;A;B)

Definition: univalence_t
univalence_t{i:l}(G) ==
  bij-contr(G.c𝕌; Σ c𝕌 Equiv(decode((q)p);decode(q)); cubical-sigma-fun(G.c𝕌;c𝕌;Equiv(decode((q)p);decode(q));UA);
  cubical-sigma-fun(G.c𝕌;c𝕌;(Path_c𝕌 (q)p q);44); 55; 66; singleton-contr(G.c𝕌;c𝕌;q)))

Definition: univalence-term
univalence-term{i:l}(G; A; B) ==  equiv-witness(UA;is-equiv-witness(G;(Path_c𝕌 B);fiber-point(33;44);55))

Definition: cubical-bool
Bool ==  encode(discr(𝔹);discrete-comp(();𝔹))

Lemma: cubical-bool_wf
Bool ∈ {() ⊢ _:c𝕌}

Lemma: csm-cubical-bool
[H:j⊢]. ∀[s:H j⟶ ()].  ((Bool)s Bool ∈ {H ⊢ _:c𝕌})

Definition: bool-negation-equiv
bool-negation-equiv(X) ==  bijection-equiv(X;𝔹;𝔹b.(¬bb);λb.(¬bb))

Lemma: bool-negation-equiv_wf
bool-negation-equiv(()) ∈ {() ⊢ _:Equiv(decode(Bool);decode(Bool))}

Definition: neg-bool-trans
neg-bool-trans() ==  PathTransport(EquivPath(();Bool;Bool;bool-negation-equiv(())))

Lemma: neg-bool-trans_wf
neg-bool-trans() ∈ {() ⊢ _:(discr(𝔹) ⟶ discr(𝔹))}

Definition: ctt-level-type
{X ⊢lvl _} ==  if (lvl =z 0) then {X ⊢ _} if (lvl =z 1) then {X ⊢_} if (lvl =z 2) then {X ⊢'' _} else {X ⊢''' _} fi 

Lemma: ctt-level-type_wf
[X:⊢''']. ∀[lvl:ℕ].  (X ⊢lvl  ∈ 𝕌{i''''})

Lemma: ctt-level-type-cumulativity
[X:⊢''']. ∀[lvl:ℕ4].  ({X ⊢lvl _} ⊆ctt-level-type{i':l}(X; lvl))

Lemma: ctt-level-type-cumulativity2
[X:⊢''']. ∀[a,b:ℕ4].  {X ⊢_} ⊆{X ⊢_} supposing a ≤ b

Lemma: ctt-level-type-subtype
[X:⊢''']. ∀[lvl:ℕ4].  ({X ⊢lvl _} ⊆{X ⊢''' _})

Lemma: ctt-level-type-subtype2
[X:⊢''']. ∀[lvl:ℕ4].  ({X ⊢ _} ⊆{X ⊢lvl _})

Definition: ctt-level-comp
Comp(X;lvl;T) ==
  if (lvl =z 0) then composition-structure{i''':l, i:l}(X; T)
  if (lvl =z 1) then composition-structure{i''':l, i':l}(X; T)
  if (lvl =z 2) then composition-structure{i''':l, i'':l}(X; T)
  else composition-structure{i''':l, i''':l}(X; T)
  fi 

Lemma: ctt-level-comp_wf
[X:⊢''']. ∀[lvl:ℕ4]. ∀[T:{X ⊢lvl _}].  (Comp(X;lvl;T) ∈ 𝕌{i'''''})

Lemma: ctt-level-comp-cumulativity
[X:⊢''']. ∀[a:ℕ4]. ∀[T:{X ⊢_}]. ∀[b:ℕ4].  Comp(X;a;T) ⊆Comp(X;b;T) supposing a ≤ b

Definition: levelsup
levelsup(x;y) ==  imax(x;y)

Lemma: levelsup_wf
[x,y:ℕ4].  (levelsup(x;y) ∈ ℕ4)

Lemma: csm-ap-type_wf-level-type
[X,Y:⊢''']. ∀[s:cube_set_map{i''':l}(X; Y)]. ∀[lvl:ℕ4]. ∀[T:{Y ⊢lvl _}].  ((T)s ∈ X ⊢lvl )

Lemma: cubical-term_wf-level-type
[X:⊢''']. ∀[lvl:ℕ4]. ∀[T:{X ⊢lvl _}].  ({X ⊢ _:T} ∈ 𝕌{i''''})

Lemma: cube-context-adjoin_wf-level-type
[X:⊢''']. ∀[lvl:ℕ4]. ∀[T:{X ⊢lvl _}].  X.T ⊢'''

Lemma: csm-ap-term_wf-level-type
[X,Y:⊢''']. ∀[s:cube_set_map{i''':l}(X; Y)]. ∀[lvl:ℕ4]. ∀[T:{Y ⊢lvl _}]. ∀[t:{Y ⊢ _:T}].  ((t)s ∈ {X ⊢ _:(T)s})

Lemma: cubical-sigma_wf-level-type
[K:⊢''']. ∀[a,b:ℕ4]. ∀[A:{K ⊢_}]. ∀[B:{K.A ⊢_}].  (Σ B ∈ K ⊢levelsup(a;b) )

Lemma: cubical-pi_wf-level-type
[K:⊢''']. ∀[a,b:ℕ4]. ∀[A:{K ⊢_}]. ∀[B:{K.A ⊢_}].  B ∈ K ⊢levelsup(a;b) )

Definition: ctt-term-meaning
cttTerm(X) ==  lvl:ℕ4 × T:{X ⊢lvl _} × {X ⊢ _:T}

Lemma: ctt-term-meaning_wf
[X:⊢''']. (cttTerm(X) ∈ 𝕌{i''''})

Definition: ctt-term-level
level(t) ==  fst(t)

Lemma: ctt-term-level_wf
[X:⊢''']. ∀[t:cttTerm(X)].  (level(t) ∈ ℕ4)

Definition: ctt-term-type
type(t) ==  fst(snd(t))

Lemma: ctt-term-type_wf
[X:⊢''']. ∀[t:cttTerm(X)].  (type(t) ∈ X ⊢level(t) )

Definition: ctt-term-type-is
type(t)=T ==  type(t) T ∈ {X ⊢''' _}

Lemma: ctt-term-type-is_wf
[X:⊢''']. ∀[t:cttTerm(X)]. ∀[T:{X ⊢''' _}].  (type(t)=T ∈ ℙ{i''''})

Definition: ctt-term-term
term(t) ==  snd(snd(t))

Lemma: ctt-term-term_wf
[X:⊢''']. ∀[t:cttTerm(X)].  (term(t) ∈ {X ⊢ _:type(t)})

Lemma: ctt-term-type-is-implies
[X:⊢''']. ∀[t:cttTerm(X)]. ∀[T:{X ⊢''' _}].  term(t) ∈ {X ⊢ _:T} supposing type(t)=T

Lemma: ctt-term-meaning-cumulativity
[X:⊢''']. (cttTerm(X) ⊆ctt-term-meaning{i':l}(X))

Lemma: ctt-term-meaning-subtype
[X,Y:⊢'''].  cttTerm(X) ⊆cttTerm(Y) supposing sub_cubical_set{i''':l}(Y; X)

Definition: mk-ctt-term-mng
mk-ctt-term-mng(lvl; T; t) ==  <lvl, T, t>

Lemma: mk-ctt-term-mng_wf
[X:⊢''']. ∀[lvl:ℕ4]. ∀[T:{X ⊢lvl _}]. ∀[t:{X ⊢ _:T}].  (mk-ctt-term-mng(lvl; T; t) ∈ cttTerm(X))

Definition: csm-ap-term-meaning
(t)s ==  let lvl,A,a in <lvl, (A)s, (a)s>

Lemma: csm-ap-term-meaning_wf
[X,Y:⊢''']. ∀[s:cube_set_map{i''':l}(X; Y)]. ∀[t:cttTerm(Y)].  ((t)s ∈ cttTerm(X))

Definition: ctt-type-meaning1
ctt-type-meaning1{i:l}(X) ==  lvl:ℕ4 × {X ⊢lvl _} × Top

Lemma: ctt-type-meaning1_wf
[X:⊢''']. (ctt-type-meaning1{i:l}(X) ∈ 𝕌{i''''})

Definition: ctt-type-meaning
cttType(X) ==  lvl:ℕ4 × T:{X ⊢lvl _} × Comp(X;lvl;T)

Lemma: ctt-type-meaning_wf
[X:⊢''']. (cttType(X) ∈ 𝕌{i'''''})

Definition: ctt-type-level
level(T) ==  fst(T)

Lemma: ctt-type-level_wf
[X:⊢''']. ∀[T:cttType(X)].  (level(T) ∈ ℕ4)

Definition: ctt-type-type
type(T) ==  fst(snd(T))

Lemma: ctt-type-type_wf
[X:⊢''']. ∀[T:cttType(X)].  (type(T) ∈ X ⊢level(T) )

Definition: ctt-type-comp
comp(T) ==  snd(snd(T))

Lemma: ctt-type-comp_wf
[X:⊢''']. ∀[T:cttType(X)].  (comp(T) ∈ Comp(X;level(T);type(T)))

Definition: mk_ctt-type-mng
cttType(levl= lvltype= Tcomp= cT) ==  <lvl, T, cT>

Lemma: mk_ctt-type-mng_wf
[X:⊢''']. ∀[lvl:ℕ4]. ∀[T:{X ⊢lvl _}]. ∀[cT:if (lvl =z 0) then composition-structure{i''':l, i:l}(X; T)
                                            if (lvl =z 1) then composition-structure{i''':l, i':l}(X; T)
                                            if (lvl =z 2) then composition-structure{i''':l, i'':l}(X; T)
                                            else composition-structure{i''':l, i''':l}(X; T)
                                            fi ].
  (cttType(levl= lvl
           type= T
           comp= cT) ∈ cttType(X))

Lemma: ctt-type-meaning-subtype
[X:⊢''']. (cttType(X) ⊆ctt-type-meaning1{i:l}(X))

Lemma: ctt-term-meaning-subtype2
[X:⊢''']. (cttTerm(X) ⊆ctt-type-meaning1{i:l}(X))

Lemma: cc-snd_wf-level-type
[X:⊢''']. ∀[lvl:ℕ4]. ∀[T:{X ⊢lvl _}].  (q ∈ {X.T ⊢ _:(T)p})

Definition: var-term-meaning
var-term-meaning(lvl;T) ==  mk-ctt-term-mng(lvl; (T)p; q)

Lemma: var-term-meaning_wf
[X:⊢''']. ∀[lvl:ℕ4]. ∀[T:{X ⊢lvl _}].  (var-term-meaning(lvl;T) ∈ cttTerm(X.T))

Definition: cubical_context
CubicalContext ==  X:CubicalSet''' × vs:varname() List × ({v:varname()| (v ∈ vs)}  ⟶ cttTerm(X))

Lemma: cubical_context_wf
CubicalContext ∈ 𝕌{i'''''}

Lemma: cubical_context_cumulativity
CubicalContext ⊆cubical_context{i':l}

Definition: in-context-dom
in-context-dom(ctxt;v) ==  let X,vs,f ctxt in (v ∈ vs)

Lemma: in-context-dom_wf
[ctxt:CubicalContext]. ∀[v:varname()].  (in-context-dom(ctxt;v) ∈ ℙ)

Definition: cubical_context_val
cubical_context_val(ctxt;v) ==  let X,vs,f ctxt in v

Lemma: cubical_context_val_wf
[ctxt:CubicalContext]. ∀[v:varname()].
  cubical_context_val(ctxt;v) ∈ cttTerm(fst(ctxt)) supposing in-context-dom(ctxt;v)

Definition: cubical-context
?CubicalContext ==  Provisional''''(CubicalContext)

Lemma: cubical-context_wf
?CubicalContext ∈ 𝕌{i'''''}

Definition: contextof
contextof(X) ==  allow(X)

Lemma: cubical-context_cumulativity
?CubicalContext ⊆?CubicalContext'

Lemma: bind-provision-cubical-context-equations
(∀[x:CubicalContext]. ∀[f:CubicalContext ⟶ ?CubicalContext].  (bind-provision(OK(x);u.f u) (f x) ∈ ?CubicalContext))
∧ (∀[m:?CubicalContext]. (bind-provision(m;u.OK(u)) m ∈ ?CubicalContext))
∧ (∀[m:?CubicalContext]. ∀[f,g:CubicalContext ⟶ ?CubicalContext].
     (bind-provision(bind-provision(m;u.f u);u.g u) bind-provision(m;u.bind-provision(f u;u.g u)) ∈ ?CubicalContext))

Lemma: trivial-same-cubical-context
X,Y:?CubicalContext.  ((Y bind-provision(X;ctxt.OK(ctxt)) ∈ ?CubicalContext)  (X Y ∈ ?CubicalContext))

Definition: context-ok
context-ok(ctxt) ==  allowed(ctxt)

Lemma: context-ok_wf
[ctxt:?CubicalContext]. (context-ok(ctxt) ∈ ℙ{i''''})

Lemma: contextof_wf
[X:?CubicalContext]. contextof(X) ∈ CubicalContext supposing context-ok(X)

Definition: context-set
context-set(ctxt) ==  fst(allow(ctxt))

Lemma: context-set_wf
[ctxt:?CubicalContext]. context-set(ctxt) ⊢''' supposing context-ok(ctxt)

Lemma: bind-provision_wf_context
[x:?CubicalContext]. ∀[f:CubicalContext ⟶ ?CubicalContext].  (bind-provision(x;t.f[t]) ∈ ?CubicalContext)

Lemma: trivial-same-context-set
X,Y:?CubicalContext.
  ((Y bind-provision(X;ctxt.OK(ctxt)) ∈ ?CubicalContext)
   context-ok(X)
   {(X Y ∈ ?CubicalContext) ∧ (context-set(X) context-set(Y) ∈ CubicalSet''')})

Definition: update-context-lvl
update-context-lvl(ctxt;lvl;T;v) ==
  let X,vs,f ctxt in 
  <X.T, [v vs], λx.if eq_var(x;v) then var-term-meaning(lvl;T) else (f x)p fi >

Lemma: update-context-lvl_wf
[ctxt:CubicalContext]. ∀[lvl:ℕ4]. ∀[T:{fst(ctxt) ⊢lvl _}]. ∀[v:varname()].
  (update-context-lvl(ctxt;lvl;T;v) ∈ CubicalContext)

Definition: update-cubical-context-I
ctxt,v:I ==  update-context-lvl(ctxt;0;𝕀;v)

Lemma: update-cubical-context-I_wf
[ctxt:CubicalContext]. ∀[v:varname()].  (ctxt,v:I ∈ CubicalContext)

Definition: update-provisional-context-I
X,v:I ==  provision(allowed(X); allow(X),v:I)

Lemma: update-provisional-context-I_wf
[ctxt:?CubicalContext]. ∀[v:varname()].  (ctxt,v:I ∈ ?CubicalContext)

Definition: update-cubical-context
ctxt; v:T ==  provision(allowed(T); let lvl,A allow(T) in update-context-lvl(ctxt;lvl;A;v))

Lemma: update-cubical-context_wf
[ctxt:CubicalContext]. ∀[v:varname()]. ∀[T:Provisional''''(lvl:ℕ4 × {fst(ctxt) ⊢lvl _})].
  (ctxt; v:T ∈ ?CubicalContext)

Definition: update-cubical-context2
ctxt; v:fst(T) ==  provision(allowed(T); let lvl,A,_ allow(T) in update-context-lvl(ctxt;lvl;A;v))

Lemma: update-cubical-context2_wf
[ctxt:CubicalContext]. ∀[v:varname()]. ∀[T:Provisional''''(ctt-type-meaning1{i:l}(fst(ctxt)))].
  (ctxt; v:fst(T) ∈ ?CubicalContext)

Definition: restrict-cubical-context
restrict-cubical-context{i:l}(ctxt;trm) ==
  provision(allowed(trm) ∧ type(allow(trm))=𝔽let X,vs,f ctxt in 
            <X, term(allow(trm)), vs, f>)

Lemma: restrict-cubical-context_wf
[ctxt:CubicalContext]. ∀[trm:Provisional''''(cttTerm(fst(ctxt)))].
  (restrict-cubical-context{i:l}(ctxt;trm) ∈ ?CubicalContext)

Definition: cubical-context-lookup
cubical-context-lookup(ctxt;v) ==  provision(in-context-dom(allow(ctxt);v); cubical_context_val(allow(ctxt);v))

Lemma: cubical-context-lookup_wf
[ctxt:?CubicalContext]. ∀[v:varname()].
  (cubical-context-lookup(ctxt;v) ∈ Provisional''''(cttTerm(context-set(ctxt))) supposing context-ok(ctxt))

Definition: update-context2
update-context2(X;v;t) ==  bind-provision(X;ctxt.ctxt; v:fst(t))

Lemma: update-context2_wf
[X:?CubicalContext]
  ∀[v:varname()]. ∀[t:Provisional''''(cttType(context-set(X)))].  (update-context2(X;v;t) ∈ ?CubicalContext) 
  supposing context-ok(X)

Lemma: context-set-update-context2
X:?CubicalContext
  (context-ok(X)
   (∀[v:varname()]. ∀[t:Provisional''''(cttType(context-set(X)))].
        ∀X':?CubicalContext
          ((X' update-context2(X;v;t) ∈ ?CubicalContext)
           {allowed(t)  (context-ok(X') ∧ (context-set(X') context-set(X).type(allow(t)) ∈ CubicalSet'''))})))

Lemma: update-context2-ok
X:?CubicalContext
  (context-ok(X)
   (∀[v:varname()]. ∀[t:Provisional''''(cttType(context-set(X)))].
        ∀X':?CubicalContext. ((X' update-context2(X;v;t) ∈ ?CubicalContext)  {allowed(t)  context-ok(X')})))

Definition: update-context3
update-context3(X;v;t) ==  bind-provision(X;ctxt.ctxt; v:fst(t))

Lemma: update-context3_wf
[X:?CubicalContext]
  ∀[v:varname()]. ∀[t:Provisional''''(cttTerm(context-set(X)))].  (update-context3(X;v;t) ∈ ?CubicalContext) 
  supposing context-ok(X)

Lemma: context-set-update-context3
X:?CubicalContext
  (context-ok(X)
   (∀[v:varname()]. ∀[t:Provisional''''(cttTerm(context-set(X)))].
        ∀X':?CubicalContext
          ((X' update-context3(X;v;t) ∈ ?CubicalContext)
           {allowed(t)  (context-ok(X') ∧ (context-set(X') context-set(X).type(allow(t)) ∈ CubicalSet'''))})))

Definition: update-context4
update-context4{i:l}(X;phi) ==  bind-provision(X;ctxt.restrict-cubical-context{i:l}(ctxt;phi))

Lemma: update-context4_wf
[X:?CubicalContext]
  ∀[phi:Provisional''''(cttTerm(context-set(X)))]. (update-context4{i:l}(X;phi) ∈ ?CubicalContext) 
  supposing context-ok(X)

Lemma: context-set-update-context4
X:?CubicalContext
  (context-ok(X)
   (∀[phi:Provisional''''(cttTerm(context-set(X)))]
        ∀X':?CubicalContext
          ((X' update-context4{i:l}(X;phi) ∈ ?CubicalContext)
           {allowed(phi)
              type(allow(phi))=𝔽
              (context-ok(X') ∧ (context-set(X') context-set(X), term(allow(phi)) ∈ CubicalSet'''))})))

Definition: ctt-tokens
ctt-tokens() ==
  ``Glue case sigma path pi decode1 decode2 decode3 encode1 encode2 encode3 pathabs pathap lambda apply pair fst ...

Lemma: ctt-tokens_wf
ctt-tokens() ∈ Atom List

Definition: ctt-op
CttOp ==
  k:{k:Atom| (k ∈ ``opid nType nterm univ``)}  × if =a "opid" then {f:Atom| (f ∈ ctt-tokens())} 
                                              if =a "nType" then Type
                                              if =a "nterm" then T:Type × T
                                              else ℕ
                                              fi 

Lemma: ctt-op_wf
CttOp ∈ 𝕌'

Definition: ctt-op-sort
ctt-op-sort(f) ==
  let k,v 
  in if =a "opid"
       then if v ∈b ``Glue decode1 decode2 decode3 case sigma path pi F`` then "fibrant"
            if =a "I" then "type"
            else "term"
            fi 
     if =a "nType" then "fibrant"
     if =a "nterm" then "term"
     else "fibrant"
     fi 

Lemma: ctt-op-sort_wf
[f:CttOp]. (ctt-op-sort(f) ∈ {x:Atom| (x ∈ ``fibrant type term``)} )

Definition: ctt-kind
ctt-kind(t) ==
  if isvarterm(t) then 0
  if ctt-op-sort(term-opr(t)) =a "fibrant" then 1
  if ctt-op-sort(term-opr(t)) =a "type" then 2
  else 0
  fi 

Lemma: ctt-kind_wf
[t:term(CttOp)]. (ctt-kind(t) ∈ ℕ)

Definition: ctt-opid-arity
ctt-opid-arity(t) ==
  if =a "inv" then [<0, 0>]
  if =a "max" then [<0, 0>; <0, 0>]
  if =a "min" then [<0, 0>; <0, 0>]
  if =a "comp" then [<0, 0>; <1, 1>; <1, 0>; <0, 0>]
  if =a "glue" then [<0, 0>; <0, 0>; <0, 0>; <0, 0>]
  if =a "unglue" then [<0, 1>; <0, 0>; <0, 1>; <0, 0>; <0, 0>]
  if =a "meet" then [<0, 0>; <0, 0>]
  if =a "join" then [<0, 0>; <0, 0>]
  if =a "eq0" then [<0, 0>]
  if =a "eq1" then [<0, 0>]
  if =a "0" then []
  if =a "1" then []
  if =a "fst" then [<0, 1>; <1, 1>; <0, 0>]
  if =a "snd" then [<0, 1>; <1, 1>; <0, 0>]
  if =a "pair" then [<0, 0>; <1, 1>; <0, 0>]
  if =a "lambda" then [<0, 1>; <1, 0>]
  if =a "apply" then [<0, 0>; <1, 1>; <0, 0>]
  if =a "pathabs" then [<0, 1>; <1, 0>]
  if =a "pathap" then [<0, 1>; <0, 0>; <0, 0>]
  if =a "decode1" then [<0, 0>]
  if =a "decode2" then [<0, 0>]
  if =a "decode3" then [<0, 0>]
  if =a "encode1" then [<0, 1>]
  if =a "encode2" then [<0, 1>]
  if =a "encode3" then [<0, 1>]
  if =a "I" then []
  if =a "F" then []
  if =a "pi" then [<0, 1>; <1, 1>]
  if =a "sigma" then [<0, 1>; <1, 1>]
  if =a "path" then [<0, 1>; <0, 0>; <0, 0>]
  if =a "case" then [<0, 0>; <0, 0>; <0, 1>; <0, 1>]
  if =a "Glue" then [<0, 1>; <0, 0>; <0, 1>; <0, 0>]
  else []
  fi 

Lemma: ctt-opid-arity_wf
[t:Atom]. (ctt-opid-arity(t) ∈ (ℕ × ℕList)

Definition: ctt-arity
ctt-arity(x) ==
  let k,v 
  in if =a "opid" then ctt-opid-arity(v)
     if =a "nType" then []
     if =a "nterm" then []
     else []
     fi 

Lemma: ctt-arity_wf
[x:CttOp]. (ctt-arity(x) ∈ (ℕ × ℕList)

Definition: ctt-term
CttTerm ==  wfterm(CttOp;λt.ctt-kind(t);λx.ctt-arity(x))

Lemma: ctt-term_wf
CttTerm ∈ 𝕌'

Lemma: ctt-term-induction
[P:CttTerm ⟶ ℙ{i'''''}]
  ((∀v:{v:varname()| ¬(v nullvar() ∈ varname())} P[varterm(v)])
   (∀f:CttOp. ∀bts:{bts:(varname() List × CttTerm) List| 
                      (||bts|| ||ctt-arity(f)|| ∈ ℤ)
                      ∧ (∀i:ℕ||bts||
                           ((||fst(bts[i])|| (fst(ctt-arity(f)[i])) ∈ ℤ)
                           ∧ (ctt-kind(snd(bts[i])) (snd(ctt-arity(f)[i])) ∈ ℤ)))} .
        ((∀i:ℕ||bts||. P[snd(bts[i])])  P[mkwfterm(f;bts)]))
   {∀t:CttTerm. P[t]})

Definition: ctt-opr-is
ctt-opr-is(f;s) ==  let k,v in =a "opid" ∧b =a s

Lemma: ctt-opr-is_wf
[f:CttOp]. ∀[s:Atom].  (ctt-opr-is(f;s) ∈ 𝔹)

Lemma: assert-ctt-opr-is
[f:CttOp]. ∀[s:Atom].  uiff(↑ctt-opr-is(f;s);f = <"opid", s> ∈ CttOp)

Lemma: ctt-opr-is-implies
[f:CttOp]. ∀[s:Atom].  ~ <"opid", s> supposing ↑ctt-opr-is(f;s)

Definition: ctt-term-is
ctt-term-is(s;t) ==  bisvarterm(t)) ∧b let k,v term-opr(t) in =a "opid" ∧b =a s

Lemma: ctt-term-is_wf
[s:Atom]. ∀[t:CttTerm].  (ctt-term-is(s;t) ∈ 𝔹)

Definition: ctt-eq
ctt-eq{i:l}(a;b) ==  alpha-eq-terms(CttOp;a;b)

Lemma: ctt-eq_wf
[a,b:CttTerm].  (ctt-eq{i:l}(a;b) ∈ ℙ')

Lemma: assert-ctt-term-is
s:Atom. ∀t:CttTerm.
  ((↑ctt-term-is(s;t))
   {(¬↑isvarterm(t))
     ∧ (term-opr(t) = <"opid", s> ∈ CttOp)
     ∧ (||term-bts(t)|| ||ctt-opid-arity(s)|| ∈ ℤ)
     ∧ (∀i:ℕ||term-bts(t)||
          ((||fst(term-bts(t)[i])|| (fst(ctt-opid-arity(s)[i])) ∈ ℤ)
          ∧ (ctt-kind(snd(term-bts(t)[i])) (snd(ctt-opid-arity(s)[i])) ∈ ℤ)))
     ∧ (t mkwfterm(term-opr(t);term-bts(t)))})

Lemma: term-accum_wf_ctt1
[R:?CubicalContext ⟶ CttTerm ⟶ ℙ{i'''''}]. ∀[Q:?CubicalContext
                                                  ⟶ f:CttOp
                                                  ⟶ vs:(varname() List)
                                                  ⟶ {L:(t:CttTerm × p:?CubicalContext × R[p;t]) List| 
                                                      ||L|| < ||ctt-arity(f)||
                                                      ∧ (||vs|| (fst(ctt-arity(f)[||L||])) ∈ ℤ)
                                                      ∧ (∀i:ℕ||L||
                                                           (ctt-kind(fst(L[i])) (snd(ctt-arity(f)[i])) ∈ ℤ))} 
                                                  ⟶ ?CubicalContext]. ∀[varcase:p:?CubicalContext
                                                                                 ⟶ v:{v:varname()| 
                                                                                       ¬(v nullvar() ∈ varname())} 
                                                                                 ⟶ R[p;varterm(v)]].
[mktermcase:p:?CubicalContext
             ⟶ f:CttOp
             ⟶ bts:wf-bound-terms(CttOp;λt.ctt-kind(t);λf.ctt-arity(f);f)
             ⟶ L:{L:(t:CttTerm × p:?CubicalContext × R[p;t]) List| 
                   (||L|| ||bts|| ∈ ℤ)
                   ∧ (∀i:ℕ||L||. ((fst(L[i])) (snd(bts[i])) ∈ term(CttOp)))
                   ∧ (∀i:ℕ||L||. ((fst(snd(L[i]))) Q[p;f;fst(bts[i]);firstn(i;L)] ∈ ?CubicalContext))} 
             ⟶ R[p;mkwfterm(f;bts)]]. ∀[t:CttTerm]. ∀[p:?CubicalContext].
  (term-accum(t with p)
   p,f,vs,tr.Q[p;f;vs;tr]
   varterm(x) with  varcase[p;x]
   mkterm(f,bts) with  trs.mktermcase[p;f;bts;trs] ∈ R[p;t])

Definition: ctt-meaning-type
ctt-meaning-type{i:l}(X;t) ==
  if isvarterm(t) then cttTerm(X)
  if ctt-op-sort(term-opr(t)) =a "fibrant" then cttType(X)
  if ctt-op-sort(term-opr(t)) =a "type" then lvl:ℕ4 × {X ⊢lvl _}
  else cttTerm(X)
  fi 

Lemma: ctt-meaning-type_wf
[X:⊢''']. ∀[t:CttTerm].  (ctt-meaning-type{i:l}(X;t) ∈ 𝕌{i'''''})

Definition: ctt_meaning
[[ctxt;t]] ==  Provisional''''(ctt-meaning-type{i:l}(context-set(ctxt);t)) supposing context-ok(ctxt)

Lemma: ctt_meaning_wf
[ctxt:?CubicalContext]. ∀[t:CttTerm].  ([[ctxt;t]] ∈ 𝕌{i'''''})

Lemma: ctt_meaning_functionality
[ctxt:?CubicalContext]. ∀[t,t':CttTerm].  [[ctxt;t]] [[ctxt;t']] ∈ 𝕌{i'''''} supposing ctt-eq{i:l}(t;t')

Definition: ctt-meaning-triple
CttMeaningTriple ==  X:?CubicalContext × bt:varname() List × CttTerm × [[X;snd(bt)]]

Lemma: ctt-meaning-triple_wf
CttMeaningTriple ∈ 𝕌{i'''''}

Definition: ctt-is-term
ctt-is-term(t) ==  isvarterm(t) ∨bctt-op-sort(term-opr(t)) =a "term"

Lemma: ctt-is-term_wf
[t:term(CttOp)]. (ctt-is-term(t) ∈ 𝔹)

Lemma: assert-ctt-is-term
[t:CttTerm]
  ∀X:?CubicalContext. [[X;t]] ⊆Provisional''''(cttTerm(context-set(X))) supposing context-ok(X) 
  supposing ↑ctt-is-term(t)

Lemma: ctt-kind-0
[t:CttTerm]
  ∀X:?CubicalContext. [[X;t]] ⊆Provisional''''(cttTerm(context-set(X))) supposing context-ok(X) 
  supposing ctt-kind(t) 0 ∈ ℤ

Definition: ctt-is-fibrant
ctt-is-fibrant(t) ==  bisvarterm(t)) ∧b ctt-op-sort(term-opr(t)) =a "fibrant"

Lemma: ctt-is-fibrant_wf
[t:term(CttOp)]. (ctt-is-fibrant(t) ∈ 𝔹)

Lemma: assert-ctt-is-fibrant
[t:CttTerm]
  ∀X:?CubicalContext. [[X;t]] ⊆Provisional''''(cttType(context-set(X))) supposing context-ok(X) 
  supposing ↑ctt-is-fibrant(t)

Lemma: ctt-kind-1
[t:CttTerm]
  ∀X:?CubicalContext. [[X;t]] ⊆Provisional''''(cttType(context-set(X))) supposing context-ok(X) 
  supposing ctt-kind(t) 1 ∈ ℤ

Definition: ctt-subterm-context
ctt-subterm-context{i:l}(ctxt;f;n;vs;m) ==
  if (ctt-opr-is(f;"pi") ∨bctt-opr-is(f;"sigma") ∨bctt-opr-is(f;"fst") ∨bctt-opr-is(f;"snd") ∨bctt-opr-is(f;"lambda"))
     ∧b (n =z 1)
    then ctxt; hd(vs):fst(m)
  if ctt-opr-is(f;"case") ∧b ((n =z 2) ∨b(n =z 3)) then restrict-cubical-context{i:l}(ctxt;m)
  if ctt-opr-is(f;"Glue") then if (n =z 2) ∨b(n =z 3) then restrict-cubical-context{i:l}(ctxt;m) else OK(ctxt) fi 
  if (ctt-opr-is(f;"pair") ∨bctt-opr-is(f;"apply")) ∧b (n =z 1) then ctxt; hd(vs):fst(m)
  if ctt-opr-is(f;"pathabs") ∧b (n =z 1) then ctxt; hd(vs):OK(<0, 𝕀>)
  if ctt-opr-is(f;"comp")
    then if (n =z 1) then ctxt; hd(vs):OK(<0, 𝕀>)
         if (n =z 2) then restrict-cubical-context{i:l}(ctxt;m),hd(vs):I
         else OK(ctxt)
         fi 
  if ctt-opr-is(f;"glue") ∧b ((n =z 2) ∨b(n =z 3)) then restrict-cubical-context{i:l}(ctxt;m)
  if ctt-opr-is(f;"unglue") ∧b ((n =z 2) ∨b(n =z 3)) then restrict-cubical-context{i:l}(ctxt;m)
  else OK(ctxt)
  fi 

Lemma: ctt-subterm-context_wf
[ctxt:CubicalContext]. ∀[f:CttOp]. ∀[n:ℕ]. ∀[vs:varname() List].
[m:Provisional''''(cttTerm(fst(ctxt))) 
    supposing ↑(((ctt-opr-is(f;"Glue") ∨bctt-opr-is(f;"case") ∨bctt-opr-is(f;"unglue")) ∧b ((n =z 2) ∨b(n =z 3)))
    ∨b(ctt-opr-is(f;"comp") ∧b (n =z 2))
    ∨b(ctt-opr-is(f;"glue") ∧b ((n =z 2) ∨b(n =z 3) ∨b(n =z 4)))) ⋂ Provisional''''(ctt-type-meaning1{i:l}(fst(ctxt))) 
                                                                    supposing ↑((ctt-opr-is(f;"pi")
                                                                    ∨bctt-opr-is(f;"sigma")
                                                                    ∨bctt-opr-is(f;"lambda")
                                                                    ∨bctt-opr-is(f;"apply")
                                                                    ∨bctt-opr-is(f;"pair")
                                                                    ∨bctt-opr-is(f;"fst")
                                                                    ∨bctt-opr-is(f;"snd"))
                                                                    ∧b (n =z 1))].
  ctt-subterm-context{i:l}(ctxt;f;n;vs;m) ∈ ?CubicalContext 
  supposing (↑(((ctt-opr-is(f;"pi")
  ∨bctt-opr-is(f;"sigma")
  ∨bctt-opr-is(f;"lambda")
  ∨bctt-opr-is(f;"apply")
  ∨bctt-opr-is(f;"pair")
  ∨bctt-opr-is(f;"fst")
  ∨bctt-opr-is(f;"snd")
  ∨bctt-opr-is(f;"pathabs")
  ∨bctt-opr-is(f;"comp"))
  ∧b (n =z 1))
  ∨b(ctt-opr-is(f;"comp") ∧b (n =z 2))))
   0 < ||vs||

Definition: provisional-subterm-context
provisional-subterm-context{i:l}(X;f;vs;L) ==
  bind-provision(X;ctxt.ctt-subterm-context{i:l}(ctxt;f;||L||;vs;if (ctt-opr-is(f;"Glue")
                                                                    ∧b ((||L|| =z 2) ∨b(||L|| =z 3)))
                                                                    ∨b(ctt-opr-is(f;"case") ∧b (||L|| =z 3))
                                                                    ∨b(ctt-opr-is(f;"unglue")
                                                                      ∧b ((||L|| =z 2) ∨b(||L|| =z 3)))
  then snd(snd(L[1]))
  else snd(snd(L[0]))
  fi ))

Lemma: provisional-subterm-context_wf
[X:?CubicalContext]. ∀[f:CttOp]. ∀[vs:varname() List]. ∀[L:{L:(a:?CubicalContext
                                                             × bt:varname() List × CttTerm
                                                             × [[a;snd(bt)]]) List| 
                                                             ||L|| < ||ctt-arity(f)||
                                                             ∧ (||vs|| (fst(ctt-arity(f)[||L||])) ∈ ℤ)
                                                             ∧ (∀i:ℕ||L||
                                                                  (ctt-kind(snd(fst(snd(L[i]))))
                                                                  (snd(ctt-arity(f)[i]))
                                                                  ∈ ℤ))
                                                             ∧ ((||L|| 1 ∈ ℤ)
                                                                ((↑ctt-opr-is(f;"pi"))
                                                                  ∨ (↑ctt-opr-is(f;"sigma"))
                                                                  ∨ (↑ctt-opr-is(f;"lambda"))
                                                                  ∨ (↑ctt-opr-is(f;"apply"))
                                                                  ∨ (↑ctt-opr-is(f;"pair"))
                                                                  ∨ (↑ctt-opr-is(f;"fst"))
                                                                  ∨ (↑ctt-opr-is(f;"snd")))
                                                                ((fst(L[0])) X ∈ ?CubicalContext))
                                                             ∧ ((||L|| 2 ∈ ℤ)
                                                                ((↑ctt-opr-is(f;"Glue"))
                                                                  ∨ (↑ctt-opr-is(f;"case"))
                                                                  ∨ (↑ctt-opr-is(f;"comp"))
                                                                  ∨ (↑ctt-opr-is(f;"unglue"))
                                                                  ∨ (↑ctt-opr-is(f;"glue")))
                                                                ((fst(L[0])) X ∈ ?CubicalContext))
                                                             ∧ ((||L|| 2 ∈ ℤ)
                                                                ((↑ctt-opr-is(f;"Glue")) ∨ (↑ctt-opr-is(f;"unglue")))
                                                                ((fst(L[1])) X ∈ ?CubicalContext))
                                                             ∧ ((||L|| 3 ∈ ℤ)
                                                                ((↑ctt-opr-is(f;"Glue"))
                                                                  ∨ (↑ctt-opr-is(f;"case"))
                                                                  ∨ (↑ctt-opr-is(f;"unglue")))
                                                                ((fst(L[1])) X ∈ ?CubicalContext))
                                                             ∧ ((||L|| 3 ∈ ℤ)
                                                                (↑ctt-opr-is(f;"glue"))
                                                                ((fst(L[0])) X ∈ ?CubicalContext))
                                                             ∧ ((||L|| 4 ∈ ℤ)
                                                                (↑ctt-opr-is(f;"glue"))
                                                                ((fst(L[0])) X ∈ ?CubicalContext))} ].
  (provisional-subterm-context{i:l}(X;f;vs;L) ∈ ?CubicalContext)

Definition: ctt-binder-context
ctt-binder-context ==
  λX,vs,f,trs,bt. if ||trs|| <||ctt-arity(f)||
                     ∧b (||fst(bt)|| =z fst(ctt-arity(f)[||trs||]))
                     ∧b bdd-all(||trs||;i.(ctt-kind(snd(fst(snd(trs[i])))) =z snd(ctt-arity(f)[i])))
                 then provisional-subterm-context{i:l}(X;f;fst(bt);trs)
                 else X
                 fi 

Lemma: ctt-binder-context_wf
ctt-binder-context ∈ ?CubicalContext
⟶ (varname() List)
⟶ CttOp
⟶ very-dep-fun(?CubicalContext;varname() List × CttTerm;a,bt.[[a;snd(bt)]])

Lemma: wfterm-accum_wf_ctt1
[m:X:?CubicalContext
    ⟶ vs:(varname() List)
    ⟶ f:CttOp
    ⟶ L:{L:CttMeaningTriple List| 
          vdf-eq(?CubicalContext;ctt-binder-context vs f;L)
          ∧ (↑wf-term(λf.ctt-arity(f);λt.ctt-kind(t);mkterm(f;map(λx.(fst(snd(x)));L))))} 
    ⟶ [[X;mkterm(f;map(λx.(fst(snd(x)));L))]]]. ∀[varcase:X:?CubicalContext
                                                           ⟶ vs:(varname() List)
                                                           ⟶ v:{v:varname()| ¬(v nullvar() ∈ varname())} 
                                                           ⟶ [[X;varterm(v)]]]. ∀[X:?CubicalContext]. ∀[t:CttTerm].
  (wfterm-accum(X;t)
   p,vs,v.varcase[p;vs;v]
   prm,vs,f,L.m[prm;vs;f;L]
   p0,ws,op,sofar,bt.ctt-binder-context[p0;ws;op;sofar;bt] ∈ [[X;t]])

Definition: eq0-meaning
eq0-meaning{i:l}(Z) ==
  let X,t,m in 
  provision(allowed(m) ∧ type(allow(m))=𝕀mk-ctt-term-mng(level(allow(m)); 𝔽(term(allow(m))=0)))

Definition: eq1-meaning
eq1-meaning{i:l}(Z) ==
  let X,t,m in 
  provision(allowed(m) ∧ type(allow(m))=𝕀mk-ctt-term-mng(level(allow(m)); 𝔽(term(allow(m))=1)))

Definition: meet-meaning
meet-meaning{i:l}(A; B) ==
  let Xa,a,ma in 
  let Xb,b,mb in 
  provision(allowed(ma) ∧ allowed(mb) ∧ type(allow(ma))=𝔽 ∧ type(allow(mb))=𝔽;
            mk-ctt-term-mng(levelsup(level(allow(ma));level(allow(mb))); 𝔽(term(allow(ma)) ∧ term(allow(mb)))))

Definition: join-meaning
join-meaning{i:l}(A; B) ==
  let Xa,a,ma in 
  let Xb,b,mb in 
  provision(allowed(ma) ∧ allowed(mb) ∧ type(allow(ma))=𝔽 ∧ type(allow(mb))=𝔽;
            mk-ctt-term-mng(levelsup(level(allow(ma));level(allow(mb))); 𝔽(term(allow(ma)) ∨ term(allow(mb)))))

Definition: max-meaning
max-meaning{i:l}(A; B) ==
  let Xa,a,ma in 
  let Xb,b,mb in 
  provision(allowed(ma) ∧ allowed(mb) ∧ type(allow(ma))=𝕀 ∧ type(allow(mb))=𝕀;
            mk-ctt-term-mng(levelsup(level(allow(ma));level(allow(mb))); 𝕀term(allow(ma)) ∨ term(allow(mb))))

Definition: min-meaning
min-meaning{i:l}(A; B) ==
  let Xa,a,ma in 
  let Xb,b,mb in 
  provision(allowed(ma) ∧ allowed(mb) ∧ type(allow(ma))=𝕀 ∧ type(allow(mb))=𝕀;
            mk-ctt-term-mng(levelsup(level(allow(ma));level(allow(mb))); 𝕀term(allow(ma)) ∧ term(allow(mb))))

Definition: inv-meaning
inv-meaning{i:l}(A) ==
  let Xa,a,ma in 
  provision(allowed(ma) ∧ type(allow(ma))=𝕀mk-ctt-term-mng(level(allow(ma)); 𝕀1-(term(allow(ma)))))

Definition: pi-meaning
pi-meaning{i:l}(A; B) ==
  let Xa,a,ma in 
  let Xb,b,mb in 
  provision(allowed(ma) ∧ allowed(mb); cttType(levl= levelsup(level(allow(ma));level(allow(mb)))
                                               type= Πtype(allow(ma)) type(allow(mb))
                                               comp= pi_comp(context-set(Xa);type(allow(ma));comp(allow(ma));
                                                             comp(allow(mb)))))

Definition: sigma-meaning
sigma-meaning{i:l}(A; B) ==
  let Xa,a,ma in 
  let Xb,b,mb in 
  provision(allowed(ma) ∧ allowed(mb); cttType(levl= levelsup(level(allow(ma));level(allow(mb)))
                                               type= Σ type(allow(ma)) type(allow(mb))
                                               comp= sigma_comp(comp(allow(ma));comp(allow(mb)))))

Definition: path-meaning
path-meaning{i:l}(A; B; C) ==
  let Xa,a,ma in 
  let Xb,b,mb in 
  let Xc,c,mc in 
  provision(allowed(ma) ∧ allowed(mb) ∧ allowed(mc) ∧ type(allow(mb))=type(allow(ma)) ∧ type(allow(mc))=type(allow(ma));
            cttType(levl= level(allow(ma))
                    type= (Path_type(allow(ma)) term(allow(mb)) term(allow(mc)))
                    comp= path_comp(context-set(Xa);type(allow(ma));term(allow(mb));term(allow(mc));
                                    comp(allow(ma)))))

Definition: pathapp-meaning
pathapp-meaning{i:l}(A; B; C) ==
  let Xa,a,ma in 
  let Xb,b,mb in 
  let Xc,c,mc in 
  provision(allowed(ma)
            ∧ allowed(mb)
            ∧ allowed(mc)
            ∧ type(allow(mc))=𝕀
            ∧ ({context-set(Xa) ⊢ _:type(allow(mb))} ⊆{context-set(Xa) ⊢ _:Path(type(allow(ma)))});
            mk-ctt-term-mng(level(allow(ma)); type(allow(ma)); term(allow(mb)) term(allow(mc))))

Definition: pathabs-meaning
pathabs-meaning{i:l}(A; B) ==
  let Xa,a,ma in 
  let Xb,b,mb in 
  provision(allowed(ma) ∧ allowed(mb) ∧ type(allow(mb))=(type(allow(ma)))p; let term(allow(mb)) in
                                                                                mk-ctt-term-mng(level(allow(ma));
                                                                                                (Path_type(allow(ma))
                                                                                                      (b)[0(𝕀)]
                                                                                                      (b)[1(𝕀)]); <>(b))\000C)

Definition: apply-meaning
apply-meaning{i:l}(A; B; C) ==
  let Xa,a,ma in 
  let Xb,b,mb in 
  let Xc,c,mc in 
  provision(allowed(ma) ∧ allowed(mb) ∧ allowed(mc) ∧ type(allow(mc))=Πtype(allow(ma)) type(allow(mb));
            mk-ctt-term-mng(level(allow(mb)); (type(allow(mb)))[term(allow(ma))];
                            cubical-apply(term(allow(mc));term(allow(ma)))))

Definition: lambda-meaning
lambda-meaning{i:l}(A; B) ==
  let Xa,dXa,ma in 
  let Xb,b,mb in 
  provision(allowed(ma) ∧ allowed(mb); mk-ctt-term-mng(levelsup(level(allow(ma));level(allow(mb)));
                                                       Πtype(allow(ma)) type(allow(mb)); term(allow(mb)))))

Definition: pair-meaning
pair-meaning{i:l}(A; B; C) ==
  let Xa,a,ma in 
  let Xb,b,mb in 
  let Xc,c,mc in 
  provision(allowed(ma) ∧ allowed(mb) ∧ allowed(mc) ∧ type(allow(mc))=(type(allow(mb)))[term(allow(ma))];
            mk-ctt-term-mng(levelsup(level(allow(ma));level(allow(mb))); Σ type(allow(ma)) type(allow(mb));
                            cubical-pair(term(allow(ma));term(allow(mc)))))

Definition: fst-meaning
fst-meaning{i:l}(A; B; C) ==
  let Xa,a,ma in 
  let Xb,b,mb in 
  let Xc,c,mc in 
  provision(allowed(ma) ∧ allowed(mb) ∧ allowed(mc) ∧ type(allow(mc))=Σ type(allow(ma)) type(allow(mb));
            mk-ctt-term-mng(level(allow(ma)); type(allow(ma)); term(allow(mc)).1))

Definition: snd-meaning
snd-meaning{i:l}(A; B; C) ==
  let Xa,a,ma in 
  let Xb,b,mb in 
  let Xc,c,mc in 
  provision(allowed(ma) ∧ allowed(mb) ∧ allowed(mc) ∧ type(allow(mc))=Σ type(allow(ma)) type(allow(mb));
            mk-ctt-term-mng(level(allow(mb)); (type(allow(mb)))[term(allow(mc)).1]; term(allow(mc)).2))

Definition: comp-meaning
comp-meaning{i:l}(A; B; C; D) ==
  let Xa,a,ma in 
  let Xb,b,mb in 
  let Xc,c,mc in 
  let Xd,d,md in 
  provision(allowed(ma)
            ∧ type(allow(ma))=𝔽
            ∧ allowed(mb)
            ∧ allowed(mc)
            ∧ allowed(md)
            ∧ type(allow(mc))=type(allow(mb))
            ∧ type(allow(md))=(type(allow(mb)))[0(𝕀)]
            ∧ ((term(allow(mc)))[0(𝕀)]
              term(allow(md))
              ∈ {context-set(Xa), term(allow(ma)) ⊢ _:(type(allow(mb)))[0(𝕀)]});
            mk-ctt-term-mng(level(allow(mb)); (type(allow(mb)))[1(𝕀)];
                            comp comp(allow(mb)) [term(allow(ma)) ⊢→ term(allow(mc))] term(allow(md))))

Definition: case-meaning
case-meaning{i:l}(A; B; C; D) ==
  let Xa,a,ma in 
  let Xb,b,mb in 
  let Xc,c,mc in 
  let Xd,d,md in 
  provision((allowed(ma) ∧ allowed(mb))
            ∧ (type(allow(ma))=𝔽
              ∧ type(allow(mb))=𝔽
              ∧ context-set(Xa) ⊢ ((term(allow(ma)) ∧ term(allow(mb)))  0(𝔽))
              ∧ context-set(Xa) ⊢ (1(𝔽 (term(allow(ma)) ∨ term(allow(mb)))))
            ∧ allowed(mc)
            ∧ allowed(md); cttType(levl= levelsup(level(allow(mc));level(allow(md)))
                                   type= (if term(allow(ma)) then type(allow(mc)) else type(allow(md)))
                                   comp= case-type-comp(context-set(Xa);
                                                        term(allow(ma));
                                                        term(allow(mb));
                                                        type(allow(mc));
                                                        type(allow(md));
                                                        comp(allow(mc));
                                                        comp(allow(md)))))

Definition: Glue-meaning
Glue-meaning{i:l}(A; B; C; D) ==
  let Xa,a,ma in 
  let Xb,b,mb in 
  let Xc,c,mc in 
  let Xd,d,md in 
  provision((allowed(ma) ∧ allowed(mb))
            ∧ type(allow(mb))=𝔽
            ∧ (allowed(mc) ∧ allowed(md))
            ∧ let type(allow(ma)) in
               let phi term(allow(mb)) in
               let type(allow(mc)) in
               type(allow(md))=Equiv(T;A); let type(allow(ma)) in
                                            let cA comp(allow(ma)) in
                                            let phi term(allow(mb)) in
                                            let type(allow(mc)) in
                                            let cT comp(allow(mc)) in
                                            let term(allow(md)) in
                                            cttType(levl= levelsup(level(allow(ma));level(allow(mc)))
                                                    type= gluetype(context-set(Xa);A;phi;T;f)
                                                    comp= comp(Glue [phi ⊢→ (T, f)] A) ))

Definition: glue-meaning
glue-meaning{i:l}(A; B; C; D) ==
  let Xa,a,ma in 
  let Xb,b,mb in 
  let Xc,c,mc in 
  let Xd,d,md in 
  provision((allowed(ma) ∧ allowed(mb))
            ∧ type(allow(ma))=𝔽
            ∧ (allowed(mc) ∧ allowed(md))
            ∧ let phi term(allow(ma)) in
                  type(allow(md))=(type(allow(mc)) ⟶ type(allow(mb)))
                  ∧ (app(term(allow(md)); term(allow(mc)))
                    term(allow(mb))
                    ∈ {context-set(Xa), phi ⊢ _:type(allow(mb))});
            let phi term(allow(ma)) in
                mk-ctt-term-mng(levelsup(level(allow(mb));level(allow(mc)));
                                Glue [phi ⊢→ (type(allow(mc));term(allow(md)))] type(allow(mb));
                                glue [phi ⊢→ term(allow(mc))] term(allow(mb))))

Definition: unglue-meaning
unglue-meaning{i:l}(A; B; C; D; E) ==
  let Xa,a,ma in 
  let Xb,b,mb in 
  let Xc,c,mc in 
  let Xd,d,md in 
  let Xe,e,me in 
  provision((allowed(ma) ∧ allowed(mb))
            ∧ type(allow(mb))=𝔽
            ∧ (allowed(mc) ∧ allowed(md) ∧ allowed(me))
            ∧ let type(allow(ma)) in
               let phi term(allow(mb)) in
               let type(allow(mc)) in
               let fst(allow(me)) in
               type(allow(md))=(T ⟶ A) ∧ type(allow(me))=Glue [phi ⊢→ (T;term(allow(md)))] A;
            mk-ctt-term-mng(level(allow(ma)); type(allow(ma)); unglue(term(allow(me)))))

Definition: decode-meaning
decode-meaning{i:l}(n; A) ==
  let Xa,a,ma in 
  provision(allowed(ma) ∧ type(allow(ma))=if (n =z 0) then c𝕌 if (n =z 1) then c𝕌else c𝕌'' fi ;
            let term(allow(ma)) in
                cttType(levl= n
                        type= decode(t)
                        comp= CompFun(t)))

Definition: encode-meaning
encode-meaning{i:l}(n; A) ==
  let Xa,a,ma in 
  provision(allowed(ma) ∧ level(allow(ma)) < n; mk-ctt-term-mng(n; if (n =z 1) then c𝕌
                                                                if (n =z 2) then c𝕌'
                                                                else c𝕌''
                                                                fi encode(type(allow(ma));
                                                                            cfun-to-cop(context-set(Xa);type(allow(ma))
                                                                                ;comp(allow(ma))))))

Definition: univ-meaning
univ-meaning{i:l}(n) ==
  OK(if (n =z 0) then cttType(levl= 1type= c𝕌comp= compU())
  if (n =z 1) then cttType(levl= 2type= c𝕌'comp= compU())
  else cttType(levl= 3
               type= c𝕌''
               comp= compU())
  fi )

Definition: ctt-op-meaning
ctt-op-meaning{i:l}(X; vs; f; L) ==
  let k,v 
  in if =a "nType" then OK(<0, discr(v), discrete_comp(context-set(X);v)>)
     if =a "nterm" then OK(let T,t in mk-ctt-term-mng(0; discr(T); discr(t)))
     if =a "opid"
       then if =a "0" then OK(mk-ctt-term-mng(0; 𝕀0(𝕀)))
            if =a "1" then OK(mk-ctt-term-mng(0; 𝕀1(𝕀)))
            if =a "eq0" then eq0-meaning{i:l}(L[0])
            if =a "eq1" then eq1-meaning{i:l}(L[0])
            if =a "meet" then meet-meaning{i:l}(L[0]; L[1])
            if =a "join" then join-meaning{i:l}(L[0]; L[1])
            if =a "max" then max-meaning{i:l}(L[0]; L[1])
            if =a "min" then min-meaning{i:l}(L[0]; L[1])
            if =a "I" then OK(<0, 𝕀>)
            if =a "F" then OK(cttType(levl= 0type= 𝔽comp= face-comp()))
            if =a "inv" then inv-meaning{i:l}(L[0])
            if =a "sigma" then sigma-meaning{i:l}(L[0]; L[1])
            if =a "pi" then pi-meaning{i:l}(L[0]; L[1])
            if =a "Glue" then Glue-meaning{i:l}(L[0]; L[1]; L[2]; L[3])
            if =a "case" then case-meaning{i:l}(L[0]; L[1]; L[2]; L[3])
            if =a "path" then path-meaning{i:l}(L[0]; L[1]; L[2])
            if =a "decode1" then decode-meaning{i:l}(0; L[0])
            if =a "decode2" then decode-meaning{i:l}(1; L[0])
            if =a "decode3" then decode-meaning{i:l}(2; L[0])
            if =a "encode1" then encode-meaning{i:l}(1; L[0])
            if =a "encode2" then encode-meaning{i:l}(2; L[0])
            if =a "encode3" then encode-meaning{i:l}(3; L[0])
            if =a "pathabs" then pathabs-meaning{i:l}(L[0]; L[1])
            if =a "pathap" then pathapp-meaning{i:l}(L[0]; L[1]; L[2])
            if =a "lambda" then lambda-meaning{i:l}(L[0]; L[1])
            if =a "apply" then apply-meaning{i:l}(L[0]; L[1]; L[2])
            if =a "pair" then pair-meaning{i:l}(L[0]; L[1]; L[2])
            if =a "fst" then fst-meaning{i:l}(L[0]; L[1]; L[2])
            if =a "snd" then snd-meaning{i:l}(L[0]; L[1]; L[2])
            if =a "glue" then glue-meaning{i:l}(L[0]; L[1]; L[2]; L[3])
            if =a "unglue" then unglue-meaning{i:l}(L[0]; L[1]; L[2]; L[3]; L[4])
            if =a "comp" then comp-meaning{i:l}(L[0]; L[1]; L[2]; L[3])
            else provision(False; ⋅)
            fi 
     else univ-meaning{i:l}(v)
     fi 

Lemma: ctt-op-meaning_wf
[X:?CubicalContext]. ∀[vs:varname() List]. ∀[f:CttOp].
[L:{L:CttMeaningTriple List| 
     vdf-eq(?CubicalContext;ctt-binder-context vs f;L)
     ∧ (↑wf-term(λf.ctt-arity(f);λt.ctt-kind(t);mkterm(f;map(λx.(fst(snd(x)));L))))} ].
  (ctt-op-meaning{i:l}(X; vs; f; L) ∈ [[X;mkterm(f;map(λx.(fst(snd(x)));L))]])

Definition: ctt-context-term-mng
ctt-context-term-mng{i:l}(X;t) ==
  wfterm-accum(X;t)
  ctxt,vs,v.cubical-context-lookup(ctxt;v)
  ctxt,vs,f,L.ctt-op-meaning{i:l}(ctxt; vs; f; L)
  p0,ws,op,sofar,bt.ctt-binder-context[p0;ws;op;sofar;bt]

Lemma: ctt-context-term-mng_wf
[X:?CubicalContext]. ∀[t:CttTerm].  (ctt-context-term-mng{i:l}(X;t) ∈ [[X;t]])

Comment: intersection type doc

We can define cubical (higher dimensional) analogue ⋂B
 of the intersection type.⋅

Definition: cubical-isect-family
cubical-isect-family(X;A;B;I;a) ==
  {w:J:fset(ℕ) ⟶ f:J ⟶ I ⟶ (⋂u:A(f(a)). B((f(a);u)))| 
   ∀J,K:fset(ℕ). ∀f:J ⟶ I. ∀g:K ⟶ J. ∀u:A(f(a)).  ((w (f(a);u) g) (w f ⋅ g) ∈ B(g((f(a);u))))} 

Lemma: cubical-isect-family_wf
[X:⊢]. ∀[A:{X ⊢ _}]. ∀[B:{X.A ⊢ _}]. ∀[I:fset(ℕ)]. ∀[a:X(I)].  (cubical-isect-family(X;A;B;I;a) ∈ Type)

Lemma: cubical-isect-family-comp
X,Delta:⊢. ∀s:Delta ⟶ X. ∀I,J:fset(ℕ). ∀f:J ⟶ I. ∀a:Delta(I). ∀A:{X ⊢ _}. ∀B:{X.A ⊢ _}.
w:cubical-isect-family(X;A;B;I;(s)a).
  K,g. (w f ⋅ g) ∈ cubical-isect-family(X;A;B;J;(s)f(a)))

Lemma: csm-cubical-isect-family
X,Delta:⊢. ∀A:{X ⊢ _}. ∀B:{X.A ⊢ _}. ∀s:Delta ⟶ X. ∀I:fset(ℕ). ∀a:Delta(I).
  (cubical-isect-family(X;A;B;I;(s)a) cubical-isect-family(Delta;(A)s;(B)(s p;q);I;a) ∈ Type)

Definition: cubical-isect
==  <λI,a. cubical-isect-family(X;A;B;I;a), λI,J,f,a,w,K,g. (w f ⋅ g)>

Lemma: cubical-isect_wf
[X:⊢]. ∀[A:{X ⊢ _}]. ∀[B:{X.A ⊢ _}].  X ⊢ ⋂B

Lemma: csm-cubical-isect
X,Delta:⊢. ∀A:{X ⊢ _}. ∀B:{X.A ⊢ _}. ∀s:Delta ⟶ X.  ((⋂B)s = ⋂(A)s (B)(s p;q) ∈ {Delta ⊢ _})

Lemma: cubical-isect-subset-adjoin2
[X,phi,A,B,C,D:Top].  (⋂~ ⋂B)

Definition: cubical-isect-elim
cubical-isect-elim(t) ==  λI,a. (t (fst(a)) 1)

Lemma: cubical-isect-elim_wf
[X:⊢]. ∀[A:{X ⊢ _}]. ∀[B:{X.A ⊢ _}]. ∀[t:{X ⊢ _:⋂B}].  (cubical-isect-elim(t) ∈ {X.A ⊢ _:B})



Home Index