mathlib documentation

data.typevec

Tuples of types, and their categorical structure.

Features:

typevec n : n-tuples of types α ⟹ β : n-tuples of maps f ⊚ g : composition

Also, support functions for operating with n-tuples of types, such as:

append1 α β : append type β to n-tuple α to obtain an (n+1)-tuple drop α : drops the last element of an (n+1)-tuple last α : returns the last element of an (n+1)-tuple append_fun f g : appends a function g to an n-tuple of functions drop_fun f : drops the last function from an n+1-tuple last_fun f : returns the last function of a tuple.

Since e.g. append1 α.drop α.last is propositionally equal to α but not definitionally equal to it, we need support functions and lemmas to mediate between constructions.

def typevec (n : ) :
Type (u_1+1)

n-tuples of types, as a category

Equations
@[instance]

Equations
def typevec.arrow {n : } (α : typevec n) (β : typevec n) :
Type (max u_1 u_2)

arrow in the category of typevec

Equations
  • β) = Π (i : fin2 n), α iβ i
@[instance]
def typevec.arrow.inhabited {n : } (α : typevec n) (β : typevec n) [Π (i : fin2 n), inhabited (β i)] :
inhabited β)

Equations
def typevec.id {n : } {α : typevec n} :
α α

identity of arrow composition

Equations
def typevec.comp {n : } {α : typevec n} {β : typevec n} {γ : typevec n} (g : β γ) (f : α β) :
α γ

arrow composition in the category of typevec

Equations
  • g f = λ (i : fin2 n) (x : α i), g i (f i x)
@[simp]
theorem typevec.id_comp {n : } {α : typevec n} {β : typevec n} (f : α β) :

@[simp]
theorem typevec.comp_id {n : } {α : typevec n} {β : typevec n} (f : α β) :

theorem typevec.comp_assoc {n : } {α : typevec n} {β : typevec n} {γ : typevec n} {δ : typevec n} (h : γ δ) (g : β γ) (f : α β) :
(h g) f = h g f

def typevec.append1 {n : } (α : typevec n) (β : Type u_1) :
typevec (n + 1)

Support for extending a typevec by one element.

Equations
def typevec.drop {n : } (α : typevec (n + 1)) :

retain only a n-length prefix of the argument

Equations
def typevec.last {n : } (α : typevec (n + 1)) :
Type u

take the last value of a (n+1)-length vector

Equations
@[instance]
def typevec.last.inhabited {n : } (α : typevec (n + 1)) [inhabited fin2.fz)] :

Equations
theorem typevec.drop_append1 {n : } {α : typevec n} {β : Type u_1} {i : fin2 n} :
::: β).drop i = α i

@[simp]
theorem typevec.drop_append1' {n : } {α : typevec n} {β : Type u_1} :
::: β).drop = α

theorem typevec.last_append1 {n : } {α : typevec n} {β : Type u_1} :
::: β).last = β

@[simp]
theorem typevec.append1_drop_last {n : } (α : typevec (n + 1)) :
α.drop ::: α.last = α

def typevec.append1_cases {n : } {C : typevec (n + 1)Sort u} (H : Π (α : typevec n) (β : Type u_1), C ::: β)) (γ : typevec (n + 1)) :
C γ

cases on (n+1)-length vectors

Equations
@[simp]
theorem typevec.append1_cases_append1 {n : } {C : typevec (n + 1)Sort u} (H : Π (α : typevec n) (β : Type u_1), C ::: β)) (α : typevec n) (β : Type u_1) :
typevec.append1_cases H ::: β) = H α β

def typevec.split_fun {n : } {α : typevec (n + 1)} {α' : typevec (n + 1)} (f : α.drop α'.drop) (g : α.last → α'.last) :
α α'

append an arrow and a function for arbitrary source and target type vectors

Equations
def typevec.append_fun {n : } {α : typevec n} {α' : typevec n} {β : Type u_1} {β' : Type u_2} (f : α α') (g : β → β') :
α ::: β α' ::: β'

append an arrow and a function as well as their respective source and target types / typevecs

Equations
def typevec.drop_fun {n : } {α : typevec (n + 1)} {β : typevec (n + 1)} (f : α β) :
α.drop β.drop

split off the prefix of an arrow

Equations
def typevec.last_fun {n : } {α : typevec (n + 1)} {β : typevec (n + 1)} (f : α β) (a : α.last) :
β.last

split off the last function of an arrow

Equations
def typevec.nil_fun {α : typevec 0} {β : typevec 0} :
α β

arrow in the category of 0-length vectors

Equations
theorem typevec.eq_of_drop_last_eq {n : } {α : typevec (n + 1)} {β : typevec (n + 1)} {f g : α β} (h₀ : typevec.drop_fun f = typevec.drop_fun g) (h₁ : typevec.last_fun f = typevec.last_fun g) :
f = g

@[simp]
theorem typevec.drop_fun_split_fun {n : } {α : typevec (n + 1)} {α' : typevec (n + 1)} (f : α.drop α'.drop) (g : α.last → α'.last) :

def typevec.arrow.mp {n : } {α β : typevec n} (h : α = β) :
α β

turn an equality into an arrow

Equations
def typevec.arrow.mpr {n : } {α β : typevec n} (h : α = β) :
β α

turn an equality into an arrow, with reverse direction

Equations
def typevec.to_append1_drop_last {n : } {α : typevec (n + 1)} :
α α.drop ::: α.last

decompose a vector into its prefix appended with its last element

Equations
def typevec.from_append1_drop_last {n : } {α : typevec (n + 1)} :
α.drop ::: α.last α

stitch two bits of a vector back together

Equations
@[simp]
theorem typevec.last_fun_split_fun {n : } {α : typevec (n + 1)} {α' : typevec (n + 1)} (f : α.drop α'.drop) (g : α.last → α'.last) :

@[simp]
theorem typevec.drop_fun_append_fun {n : } {α : typevec n} {α' : typevec n} {β : Type u_1} {β' : Type u_2} (f : α α') (g : β → β') :

@[simp]
theorem typevec.last_fun_append_fun {n : } {α : typevec n} {α' : typevec n} {β : Type u_1} {β' : Type u_2} (f : α α') (g : β → β') :

theorem typevec.split_drop_fun_last_fun {n : } {α : typevec (n + 1)} {α' : typevec (n + 1)} (f : α α') :

theorem typevec.split_fun_inj {n : } {α : typevec (n + 1)} {α' : typevec (n + 1)} {f f' : α.drop α'.drop} {g g' : α.last → α'.last} (H : typevec.split_fun f g = typevec.split_fun f' g') :
f = f' g = g'

theorem typevec.append_fun_inj {n : } {α : typevec n} {α' : typevec n} {β : Type u_1} {β' : Type u_2} {f f' : α α'} {g g' : β → β'} (a : f ::: g = f' ::: g') :
f = f' g = g'

theorem typevec.split_fun_comp {n : } {α₀ : typevec (n + 1)} {α₁ : typevec (n + 1)} {α₂ : typevec (n + 1)} (f₀ : α₀.drop α₁.drop) (f₁ : α₁.drop α₂.drop) (g₀ : α₀.last → α₁.last) (g₁ : α₁.last → α₂.last) :
typevec.split_fun (f₁ f₀) (g₁ g₀) = typevec.split_fun f₁ g₁ typevec.split_fun f₀ g₀

theorem typevec.append_fun_comp_split_fun {n : } {α : typevec n} {γ : typevec n} {β : Type u_1} {δ : Type u_2} {ε : typevec (n + 1)} (f₀ : ε.drop α) (f₁ : α γ) (g₀ : ε.last → β) (g₁ : β → δ) :
(f₁ ::: g₁) typevec.split_fun f₀ g₀ = typevec.split_fun (f₁ f₀) (g₁ g₀)

theorem typevec.append_fun_comp {n : } {α₀ : typevec n} {α₁ : typevec n} {α₂ : typevec n} {β₀ : Type u_1} {β₁ : Type u_2} {β₂ : Type u_3} (f₀ : α₀ α₁) (f₁ : α₁ α₂) (g₀ : β₀ → β₁) (g₁ : β₁ → β₂) :
f₁ f₀ ::: g₁ g₀ = (f₁ ::: g₁) (f₀ ::: g₀)

theorem typevec.append_fun_comp' {n : } {α₀ : typevec n} {α₁ : typevec n} {α₂ : typevec n} {β₀ : Type u_1} {β₁ : Type u_2} {β₂ : Type u_3} (f₀ : α₀ α₁) (f₁ : α₁ α₂) (g₀ : β₀ → β₁) (g₁ : β₁ → β₂) :
(f₁ ::: g₁) (f₀ ::: g₀) = f₁ f₀ ::: g₁ g₀

theorem typevec.nil_fun_comp {α₀ : typevec 0} (f₀ : α₀ fin2.elim0) :

theorem typevec.append_fun_comp_id {n : } {α : typevec n} {β₀ β₁ β₂ : Type u_1} (g₀ : β₀ → β₁) (g₁ : β₁ → β₂) :
typevec.id ::: g₁ g₀ = (typevec.id ::: g₁) (typevec.id ::: g₀)

@[simp]
theorem typevec.drop_fun_comp {n : } {α₀ : typevec (n + 1)} {α₁ : typevec (n + 1)} {α₂ : typevec (n + 1)} (f₀ : α₀ α₁) (f₁ : α₁ α₂) :

@[simp]
theorem typevec.last_fun_comp {n : } {α₀ : typevec (n + 1)} {α₁ : typevec (n + 1)} {α₂ : typevec (n + 1)} (f₀ : α₀ α₁) (f₁ : α₁ α₂) :

theorem typevec.append_fun_aux {n : } {α : typevec n} {α' : typevec n} {β : Type u_1} {β' : Type u_2} (f : α ::: β α' ::: β') :

theorem typevec.append_fun_id_id {n : } {α : typevec n} {β : Type u_1} :

simp set for the manipulation of typevec and arrow expressions

def typevec.cases_nil {β : typevec 0Sort u_2} (f : β fin2.elim0) (v : typevec 0) :
β v

cases distinction for 0-length type vector

Equations
def typevec.cases_cons (n : ) {β : typevec (n + 1)Sort u_2} (f : Π (t : Type u_1) (v : typevec n), β (v ::: t)) (v : typevec (n + 1)) :
β v

cases distinction for (n+1)-length type vector

Equations
theorem typevec.cases_nil_append1 {β : typevec 0Sort u_2} (f : β fin2.elim0) :

theorem typevec.cases_cons_append1 (n : ) {β : typevec (n + 1)Sort u_2} (f : Π (t : Type u_1) (v : typevec n), β (v ::: t)) (v : typevec n) (α : Type u_1) :
typevec.cases_cons n f (v ::: α) = f α v

def typevec.typevec_cases_nil₃ {β : Π (v : typevec 0) (v' : typevec 0), v v'Sort u_3} (f : β fin2.elim0 fin2.elim0 typevec.nil_fun) (v : typevec 0) (v' : typevec 0) (fs : v v') :
β v v' fs

cases distinction for an arrow in the category of 0-length type vectors

Equations
def typevec.typevec_cases_cons₃ (n : ) {β : Π (v : typevec (n + 1)) (v' : typevec (n + 1)), v v'Sort u_3} (F : Π (t : Type u_1) (t' : Type u_2) (f : t → t') (v : typevec n) (v' : typevec n) (fs : v v'), β (v ::: t) (v' ::: t') (fs ::: f)) (v : typevec (n + 1)) (v' : typevec (n + 1)) (fs : v v') :
β v v' fs

cases distinction for an arrow in the category of (n+1)-length type vectors

Equations
def typevec.typevec_cases_nil₂ {β : fin2.elim0 fin2.elim0Sort u_3} (f : β typevec.nil_fun) (f_1 : fin2.elim0 fin2.elim0) :
β f_1

specialized cases distinction for an arrow in the category of 0-length type vectors

Equations
def typevec.typevec_cases_cons₂ (n : ) (t : Type u_1) (t' : Type u_2) (v : typevec n) (v' : typevec n) {β : v ::: t v' ::: t'Sort u_3} (F : Π (f : t → t') (fs : v v'), β (fs ::: f)) (fs : v ::: t v' ::: t') :
β fs

specialized cases distinction for an arrow in the category of (n+1)-length type vectors

Equations
theorem typevec.typevec_cases_cons₂_append_fun (n : ) (t : Type u_1) (t' : Type u_2) (v : typevec n) (v' : typevec n) {β : v ::: t v' ::: t'Sort u_3} (F : Π (f : t → t') (fs : v v'), β (fs ::: f)) (f : t → t') (fs : v v') :
typevec.typevec_cases_cons₂ n t t' v v' F (fs ::: f) = F f fs

def typevec.pred_last {n : } (α : typevec n) {β : Type u_1} (p : β → Prop) ⦃i : fin2 (n + 1) (a : ::: β) i) :
Prop

pred_last α p x predicates p of the last element of x : α.append1 β.

Equations
def typevec.rel_last {n : } (α : typevec n) {β γ : Type u_1} (r : β → γ → Prop) ⦃i : fin2 (n + 1) (a : ::: β) i) (a_1 : ::: γ) i) :
Prop

rel_last α r x y says that p the last elements of x y : α.append1 β are related by r and all the other elements are equal.

Equations
def typevec.repeat (n : ) (t : Type u_1) :

repeat n t is a n-length type vector that contains n occurences of t

Equations
def typevec.prod {n : } (α β : typevec n) :

prod α β is the pointwise product of the components of α and β

Equations
def typevec.const {β : Type u_1} (x : β) {n : } (α : typevec n) :

const x α is an arrow that ignores its source and constructs a typevec that contains nothing but x

Equations
def typevec.repeat_eq {n : } (α : typevec n) :
α α typevec.repeat n Prop

vector of equality on a product of vectors

Equations
theorem typevec.const_append1 {β : Type u_1} {γ : Type u_2} (x : γ) {n : } (α : typevec n) :
typevec.const x ::: β) = typevec.const x α ::: λ (_x : β), x

theorem typevec.eq_nil_fun {α : typevec 0} {β : typevec 0} (f : α β) :

theorem typevec.const_nil {β : Type u_1} (x : β) (α : typevec 0) :

theorem typevec.repeat_eq_append1 {β : Type u_1} {n : } (α : typevec n) :

def typevec.pred_last' {n : } (α : typevec n) {β : Type u_1} (p : β → Prop) :
α ::: β typevec.repeat (n + 1) Prop

predicate on a type vector to constrain only the last object

Equations
def typevec.rel_last' {n : } (α : typevec n) {β : Type u_1} (p : β → β → Prop) :
α ::: β α ::: β typevec.repeat (n + 1) Prop

predicate on the product of two type vectors to constrain only their last object

Equations
def typevec.curry {n : } (F : typevec (n + 1)Type u_1) (α : Type u) (β : typevec n) :
Type u_1

given F : typevec.{u} (n+1) → Type u, curry F : Type u → typevec.{u} → Type u, i.e. its first argument can be fed in separately from the rest of the vector of arguments

Equations
@[instance]
def typevec.curry.inhabited {n : } (F : typevec (n + 1)Type u_1) (α : Type u) (β : typevec n) [I : inhabited (F ::: α))] :

Equations
def typevec.drop_repeat (α : Type u_1) {n : } :

arrow to remove one element of a repeat vector

Equations
def typevec.of_repeat {α : Type u_1} {n : } {i : fin2 n} (a : typevec.repeat n α i) :
α

projection for a repeat vector

Equations
theorem typevec.const_iff_true {n : } {α : typevec n} {i : fin2 n} {x : α i} {p : Prop} :

def typevec.prod.fst {n : } {α β : typevec n} :
α β α

left projection of a prod vector

Equations
def typevec.prod.snd {n : } {α β : typevec n} :
α β β

right projection of a prod vector

Equations
def typevec.prod.diag {n : } {α : typevec n} :
α α α

introduce a product where both components are the same

Equations
def typevec.prod.mk {n : } {α β : typevec n} (i : fin2 n) (a : α i) (a_1 : β i) :
β) i

constructor for prod

Equations
def typevec.prod.map {n : } {α α' β β' : typevec n} (a : α β) (a_1 : α' β') :
α α' β β'

prod is functorial

Equations
theorem typevec.fst_prod_mk {n : } {α α' β β' : typevec n} (f : α β) (g : α' β') :

theorem typevec.snd_prod_mk {n : } {α α' β β' : typevec n} (f : α β) (g : α' β') :

theorem typevec.repeat_eq_iff_eq {n : } {α : typevec n} {i : fin2 n} {x y : α i} :

def typevec.subtype_ {n : } {α : typevec n} (p : α typevec.repeat n Prop) :

given a predicate vector p over vector α, subtype_ p is the type of vectors that contain an α that satisfies p

Equations
def typevec.to_subtype {n : } {α : typevec n} (p : α typevec.repeat n Prop) :
(λ (i : fin2 n), {x // typevec.of_repeat (p i x)}) typevec.subtype_ p

arrow that rearranges the type of subtype_ to turn a subtype of vector into a vector of subtypes

Equations
def typevec.of_subtype {n : } {α : typevec n} (p : α typevec.repeat n Prop) :
typevec.subtype_ p λ (i : fin2 n), {x // typevec.of_repeat (p i x)}

arrow that rearranges the type of subtype_ to turn a vector of subtypes into a subtype of vector

Equations
def typevec.to_subtype' {n : } {α : typevec n} (p : α α typevec.repeat n Prop) :
(λ (i : fin2 n), {x // typevec.of_repeat (p i (typevec.prod.mk i x.fst x.snd))}) typevec.subtype_ p

similar to to_subtype adapted to relations (i.e. predicate on product)

Equations
def typevec.of_subtype' {n : } {α : typevec n} (p : α α typevec.repeat n Prop) :
typevec.subtype_ p λ (i : fin2 n), {x // typevec.of_repeat (p i (typevec.prod.mk i x.fst x.snd))}

similar to of_subtype adapted to relations (i.e. predicate on product)

Equations

similar to diag but the target vector is a subtype_ guaranteeing the equality of the components

Equations
theorem typevec.append_prod_append_fun {n : } {α α' β β' : typevec n} {φ φ' ψ ψ' : Type u} {f₀ : α α'} {g₀ : β β'} {f₁ : φ → φ'} {g₁ : ψ → ψ'} :
(f₀ ⊗' g₀) ::: prod.map f₁ g₁ = (f₀ ::: f₁ ⊗' g₀ ::: g₁)

@[simp]

@[simp]
theorem typevec.drop_fun_to_subtype {n : } {α : typevec (n + 1)} (p : α typevec.repeat (n + 1) Prop) :
typevec.drop_fun (typevec.to_subtype p) = typevec.to_subtype (λ (i : fin2 n) (x : α i.fs), p i.fs x)

@[simp]
theorem typevec.last_fun_to_subtype {n : } {α : typevec (n + 1)} (p : α typevec.repeat (n + 1) Prop) :

@[simp]
theorem typevec.last_fun_of_subtype {n : } {α : typevec (n + 1)} (p : α typevec.repeat (n + 1) Prop) :

@[simp]
theorem typevec.drop_fun_rel_last {n : } {α : typevec n} {β : Type u_1} (R : β → β → Prop) :

@[simp]
theorem typevec.drop_fun_prod {n : } {α α' β β' : typevec (n + 1)} (f : α β) (f' : α' β') :

@[simp]
theorem typevec.last_fun_prod {n : } {α α' β β' : typevec (n + 1)} (f : α β) (f' : α' β') :

@[simp]

@[simp]
theorem typevec.to_subtype_of_subtype {n : } {α : typevec n} (p : α typevec.repeat n Prop) :
typevec.to_subtype p typevec.of_subtype (λ (i : fin2 n) (x : α i), p i x) = typevec.id

@[simp]
theorem typevec.to_subtype_of_subtype_assoc {n : } {α : typevec n} {β : typevec n} (p : α typevec.repeat n Prop) (f : β typevec.subtype_ p) :
typevec.to_subtype p typevec.of_subtype (λ (i : fin2 n) (x : α i), p i x) f = f